diff --git a/_CoqProject b/_CoqProject index c5bc489..cedd088 100644 --- a/_CoqProject +++ b/_CoqProject @@ -39,6 +39,7 @@ theories/type_systems/stlc_extended/logrel_sol.v theories/type_systems/systemf/lang.v theories/type_systems/systemf/notation.v theories/type_systems/systemf/types.v +theories/type_systems/systemf/types_sol.v theories/type_systems/systemf/tactics.v theories/type_systems/systemf/bigstep.v theories/type_systems/systemf/church_encodings.v @@ -47,8 +48,19 @@ theories/type_systems/systemf/logrel.v theories/type_systems/systemf/logrel_sol.v theories/type_systems/systemf/free_theorems.v theories/type_systems/systemf/binary_logrel.v +theories/type_systems/systemf/binary_logrel_sol.v theories/type_systems/systemf/existential_invariants.v +# SystemF-Mu +theories/type_systems/systemf_mu/lang.v +theories/type_systems/systemf_mu/notation.v +theories/type_systems/systemf_mu/types.v +theories/type_systems/systemf_mu/tactics.v +theories/type_systems/systemf_mu/pure.v +theories/type_systems/systemf_mu/untyped_encoding.v +theories/type_systems/systemf_mu/parallel_subst.v +theories/type_systems/systemf_mu/logrel.v + # By removing the # below, you can add the exercise sheets to make theories/type_systems/warmup/warmup.v #theories/type_systems/warmup/warmup_sol.v @@ -56,10 +68,12 @@ theories/type_systems/stlc/exercises01.v #theories/type_systems/stlc/exercises01_sol.v theories/type_systems/stlc/exercises02.v #theories/type_systems/stlc/exercises02_sol.v -#theories/type_systems/stlc/cbn_logrel.v +# theories/type_systems/stlc/cbn_logrel.v #theories/type_systems/stlc/cbn_logrel_sol.v -#theories/type_systems/systemf/exercises03.v +theories/type_systems/systemf/exercises03.v #theories/type_systems/systemf/exercises03_sol.v -#theories/type_systems/systemf/exercises04.v +theories/type_systems/systemf/exercises04.v #theories/type_systems/systemf/exercises04_sol.v theories/type_systems/systemf/exercises05.v +#theories/type_systems/systemf/exercises05_sol.v +#theories/type_systems/systemf_mu/exercises06.v diff --git a/theories/type_systems/systemf/binary_logrel_sol.v b/theories/type_systems/systemf/binary_logrel_sol.v new file mode 100644 index 0000000..8f430b6 --- /dev/null +++ b/theories/type_systems/systemf/binary_logrel_sol.v @@ -0,0 +1,1133 @@ +(* NOTE: import order matters here. + stdpp and Equations both set different default obligation tactics, + and the one from Equations is much better at solving the Equations goals. *) +From stdpp Require Import gmap base relations. +From iris Require Import prelude. +From semantics.lib Require Export facts. +From semantics.ts.systemf Require Import lang notation parallel_subst types bigstep tactics. +From Equations Require Export Equations. + +(** * Logical relation for SystemF *) + +Implicit Types + (Δ : nat) + (Γ : typing_context) + (v : val) + (α : var) + (e : expr) + (A : type). + +(** ** Definition of the logrel *) +(** + In Coq, we need to make argument why the logical relation is well-defined precise: + This holds true in particular for the mutual recursion between the value relation and the expression relation. + We therefore define a termination measure [mut_measure] that makes sure that for each recursive call, we either + - decrease the size of the type + - or switch from the expression case to the value case. + + We use the Equations package to define the logical relation, as it's tedious to make the termination + argument work with Coq's built-in support for recursive functions. + *) +Inductive val_or_expr : Type := +| inj_val : val → val → val_or_expr +| inj_expr : expr → expr → val_or_expr. + +(* The [type_size] function essentially computes the size of the "type tree". *) +(* Note that we have added some additional primitives to make our (still + simple) language more expressive. *) +Equations type_size (A : type) : nat := + type_size Int := 1; + type_size Bool := 1; + type_size Unit := 1; + type_size (A → B) := type_size A + type_size B + 1; + type_size (#α) := 1; + type_size (∀: A) := type_size A + 2; + type_size (∃: A) := type_size A + 2; + type_size (A × B) := type_size A + type_size B + 1; + type_size (A + B) := max (type_size A) (type_size B) + 1. + +(* The definition of the expression relation uses the value relation -- therefore, it needs to be larger, and we add [1]. *) +Equations mut_measure (ve : val_or_expr) (t : type) : nat := + mut_measure (inj_val _ _) t := type_size t; + mut_measure (inj_expr _ _) t := 1 + type_size t. + +(** A semantic type consists of a value-relation and a proof of closedness *) +Record sem_type := mk_ST { + sem_type_car :> val → val → Prop; + sem_type_closed_val v w : sem_type_car v w → is_closed [] (of_val v) ∧ is_closed [] (of_val w); + }. +(** Two tactics we will use later on. + [pose_sem_type P as N] defines a semantic type at name [N] with the value predicate [P]. + [specialize_sem_type S with P as N] specializes a universal quantifier over sem types in [S] with + a semantic type with predicate [P], giving it the name [N]. + *) +(* slightly complicated formulation to make the proof term [p] opaque and prevent it from polluting the context *) +Tactic Notation "pose_sem_type" uconstr(P) "as" ident(N) := + let p := fresh "__p" in + unshelve refine ((λ p, let N := (mk_ST P p) in _) _); first (simpl in p); cycle 1. +Tactic Notation "specialize_sem_type" constr(S) "with" uconstr(P) "as" ident(N) := + pose_sem_type P as N; last specialize (S N). + +(** We represent type variable assignments [δ] as lists of semantic types. + The variable #n (in De Bruijn representation) is mapped to the [n]-th element of the list. + *) +Definition tyvar_interp := nat → sem_type. +Implicit Types + (δ : tyvar_interp) + (τ : sem_type) +. + +(** The logical relation *) +Equations type_interp (c : val_or_expr) (t : type) δ : Prop by wf (mut_measure c t) := { + type_interp (inj_val v w) Int δ => + ∃ z : Z, v = #z ∧ w = #z; + type_interp (inj_val v w) Bool δ => + ∃ b : bool, v = #b ∧ w = #b; + type_interp (inj_val v w) Unit δ => + v = #LitUnit ∧ w = #LitUnit ; + type_interp (inj_val v w) (A × B) δ => + ∃ v1 v2 w1 w2 : val, v = (v1, v2)%V ∧ w = (w1, w2)%V ∧ type_interp (inj_val v1 w1) A δ ∧ type_interp (inj_val v2 w2) B δ; + type_interp (inj_val v w) (A + B) δ => + (∃ v' w' : val, v = InjLV v' ∧ w = InjLV w' ∧ type_interp (inj_val v' w') A δ) ∨ + (∃ v' w' : val, v = InjRV v' ∧ w = InjRV w' ∧ type_interp (inj_val v' w') B δ); + type_interp (inj_val v w) (A → B) δ => + ∃ x y e1 e2, v = LamV x e1 ∧ w = LamV y e2 ∧ is_closed (x :b: nil) e1 ∧ is_closed (y :b: nil) e2 ∧ + ∀ v' w', + type_interp (inj_val v' w') A δ → + type_interp (inj_expr (subst' x (of_val v') e1) (subst' y (of_val w') e2)) B δ; + (** Type variable case *) + type_interp (inj_val v w) (#α) δ => + (δ α).(sem_type_car) v w; + (** ∀ case *) + type_interp (inj_val v w) (∀: A) δ => + ∃ e1 e2, v = TLamV e1 ∧ w = TLamV e2 ∧ is_closed [] e1 ∧ is_closed [] e2 ∧ + ∀ τ, type_interp (inj_expr e1 e2) A (τ .: δ); + (** ∃ case *) + type_interp (inj_val v w) (∃: A) δ => + ∃ v' w', v = PackV v' ∧ w = PackV w' ∧ + ∃ τ : sem_type, type_interp (inj_val v' w') A (τ .: δ); + + type_interp (inj_expr e1 e2) t δ => + ∃ v1 v2, big_step e1 v1 ∧ big_step e2 v2 ∧ type_interp (inj_val v1 v2) t δ +}. +Next Obligation. repeat simp mut_measure; simp type_size; lia. Qed. +Next Obligation. repeat simp mut_measure; simp type_size; lia. Qed. +Next Obligation. repeat simp mut_measure; simp type_size; lia. Qed. +Next Obligation. + simp mut_measure. simp type_size. + destruct A; repeat simp mut_measure; repeat simp type_size; lia. +Qed. +Next Obligation. repeat simp mut_measure; simp type_size; lia. Qed. +Next Obligation. repeat simp mut_measure; simp type_size; lia. Qed. +Next Obligation. repeat simp mut_measure; simp type_size; lia. Qed. +Next Obligation. repeat simp mut_measure; simp type_size; lia. Qed. + +(** Value relation and expression relation *) +Notation sem_val_rel A δ v w := (type_interp (inj_val v w) A δ). +Notation sem_expr_rel A δ e1 e2 := (type_interp (inj_expr e1 e2) A δ). + +Notation 𝒱 A δ v w := (sem_val_rel A δ v w). +Notation ℰ A δ e1 e2 := (sem_expr_rel A δ e1 e2). + + +Lemma val_rel_is_closed v w δ A: + 𝒱 A δ v w → is_closed [] (of_val v) ∧ is_closed [] (of_val w). +Proof. + induction A as [ | | | | | A IHA | | A IH1 B IH2 | A IH1 B IH2] in v, w, δ |-*; simp type_interp. + - by eapply sem_type_closed_val. + - intros [z [-> ->]]. done. + - intros [b [-> ->]]. done. + - intros [-> ->]. done. + - intros (e1 & e2 & -> & -> & ? & ? & _). done. + - intros (v' & w' & -> & -> & (τ & Hinterp)). simpl. by eapply IHA. + - intros (x & y & e1 & e2 & -> & -> & ? & ? & _). done. + - intros (v1 & v2 & w1 & w2 & -> & -> & ? & ?). simpl. split; apply andb_True; split; naive_solver. + - intros [(v' & w' & -> & -> & ?) | (v' & w' & -> & -> & ?)]; simpl; eauto. +Qed. + +(** Interpret a syntactic type *) +Program Definition interp_type A δ : sem_type := {| + sem_type_car := fun v w => 𝒱 A δ v w; +|}. +Next Obligation. by eapply val_rel_is_closed. Qed. + +(* Semantic typing of contexts *) +Implicit Types + (θ : gmap string expr). + +(** Context relation *) +Inductive sem_context_rel (δ : tyvar_interp) : typing_context → (gmap string expr) → (gmap string expr) → Prop := + | sem_context_rel_empty : sem_context_rel δ ∅ ∅ ∅ + | sem_context_rel_insert Γ θ1 θ2 v w x A : + 𝒱 A δ v w → + sem_context_rel δ Γ θ1 θ2 → + sem_context_rel δ (<[x := A]> Γ) (<[x := of_val v]> θ1) (<[x := of_val w]> θ2). + +Notation 𝒢 := sem_context_rel. + +(** Semantic typing judgment *) +Definition sem_typed Δ Γ e1 e2 A := + is_closed (elements (dom Γ)) e1 ∧ is_closed (elements (dom Γ)) e2 ∧ + ∀ θ1 θ2 δ, 𝒢 δ Γ θ1 θ2 → ℰ A δ (subst_map θ1 e1) (subst_map θ2 e2). +Notation "'TY' Δ ; Γ ⊨ e1 ≈ e2 : A" := (sem_typed Δ Γ e1 e2 A) (at level 74, e1, e2, A at next level). + +Lemma sem_expr_rel_of_val A δ v w : + ℰ A δ (of_val v) (of_val w) → 𝒱 A δ v w. +Proof. + simp type_interp. + intros (v' & w' & ->%big_step_val & ->%big_step_val & Hv'). + apply Hv'. +Qed. + +Lemma sem_context_rel_vals {δ Γ θ1 θ2 x A} : + sem_context_rel δ Γ θ1 θ2 → + Γ !! x = Some A → + ∃ e1 e2 v1 v2, θ1 !! x = Some e1 ∧ θ2 !! x = Some e2 ∧ to_val e1 = Some v1 ∧ to_val e2 = Some v2 ∧ 𝒱 A δ v1 v2. +Proof. + induction 1 as [|Γ θ1 θ2 v w y B Hvals Hctx IH]. + - naive_solver. + - rewrite lookup_insert_Some. intros [[-> ->]|[Hne Hlook]]. + + do 4 eexists. + split; first by rewrite lookup_insert. + split; first by rewrite lookup_insert. + split; first by eapply to_of_val. + split; first by eapply to_of_val. + done. + + eapply IH in Hlook as (e1 & e2 & w1 & w2 & Hlook1 & Hlook2 & He1 & He2 & Hval). + do 4 eexists. + split; first by rewrite lookup_insert_ne. + split; first by rewrite lookup_insert_ne. + repeat split; done. +Qed. + +Lemma sem_context_rel_subset δ Γ θ1 θ2 : + 𝒢 δ Γ θ1 θ2 → dom Γ ⊆ dom θ1 ∧ dom Γ ⊆ dom θ2. +Proof. + intros Hctx. split; apply elem_of_subseteq; intros x (A & Hlook)%elem_of_dom. + all: eapply sem_context_rel_vals in Hlook as (e1 & e2 & v1 & v2 & Hlook1 & Hlook2 & Heq1 & Heq2 & Hval); last done. + all: eapply elem_of_dom; eauto. +Qed. + + +Lemma sem_context_rel_closed δ Γ θ1 θ2: + 𝒢 δ Γ θ1 θ2 → subst_is_closed [] θ1 ∧ subst_is_closed [] θ2. +Proof. + induction 1 as [ | Γ θ1 θ2 v w x A Hv Hctx [IH1 IH2]]; rewrite /subst_is_closed. + - naive_solver. + - split; intros y e; rewrite lookup_insert_Some. + all: intros [[-> <-]|[Hne Hlook]]. + + eapply val_rel_is_closed in Hv. naive_solver. + + eapply IH1; last done. + + eapply val_rel_is_closed in Hv. naive_solver. + + eapply IH2; last done. +Qed. +Lemma sem_context_rel_dom δ Γ θ1 θ2 : + 𝒢 δ Γ θ1 θ2 → dom Γ = dom θ1 /\ dom Γ = dom θ2. +Proof. + induction 1. + - by rewrite !dom_empty. + - rewrite !dom_insert. set_solver. +Qed. + + +Section boring_lemmas. + (** The lemmas in this section are all quite boring and expected statements, + but are quite technical to prove due to De Bruijn binders. + We encourage to skip over the proofs of these lemmas. + *) + + Lemma sem_val_rel_ext B δ δ' v w : + (∀ n v w, δ n v w ↔ δ' n v w) → + 𝒱 B δ v w ↔ 𝒱 B δ' v w. + Proof. + induction B in δ, δ', v, w |-*; simpl; simp type_interp; eauto; intros Hiff. + - f_equiv; intros e1. f_equiv; intros e2. do 4 f_equiv. + eapply forall_proper; intros τ. + simp type_interp. f_equiv. intros v1. f_equiv. intros v2. + do 2 f_equiv. etransitivity; last apply IHB; first done. + intros [|m] ?; simpl; eauto. + - f_equiv. intros v1. f_equiv. intros v2. do 3 f_equiv. intros τ. + etransitivity; last apply IHB; first done. + intros [|m] ?; simpl; eauto. + - do 4 (f_equiv; intros ?). + do 4 f_equiv. + eapply forall_proper; intros v1. + eapply forall_proper; intros v2. + eapply if_iff; first eauto. + simp type_interp. f_equiv. intros v3. + f_equiv. intros v4. + do 2 f_equiv. + eauto. + - do 4 (f_equiv; intros ?). + do 3 f_equiv; eauto. + - f_equiv; f_equiv; intros ?; f_equiv; intros ?. + all: do 2 f_equiv; eauto. + Qed. + + + Lemma sem_val_rel_move_ren B δ σ v w: + 𝒱 B (λ n, δ (σ n)) v w ↔ 𝒱 (rename σ B) δ v w. + Proof. + induction B in σ, δ, v, w |-*; simpl; simp type_interp; eauto. + - f_equiv; intros e1. f_equiv; intros e2. do 4 f_equiv. + eapply forall_proper; intros τ. + simp type_interp. f_equiv. intros v1. f_equiv. intros v2. + do 2 f_equiv. etransitivity; last apply IHB. + eapply sem_val_rel_ext; intros [|n] u; asimpl; done. + - f_equiv. intros v1. f_equiv. intros v2. do 3 f_equiv. intros τ. + etransitivity; last apply IHB. + eapply sem_val_rel_ext; intros [|n] u. + + simp type_interp. done. + + simpl. rewrite /up. simpl. done. + - do 4 (f_equiv; intros ?). + do 4 f_equiv. + eapply forall_proper; intros v1. + eapply forall_proper; intros v2. + eapply if_iff; first eauto. + simp type_interp. f_equiv. intros v3. + f_equiv. intros v4. + do 2 f_equiv. + eauto. + - do 4 (f_equiv; intros ?). + do 3 f_equiv; eauto. + - f_equiv; f_equiv; intros ?; f_equiv; intros ?. + all: do 2 f_equiv; eauto. + Qed. + + + Lemma sem_val_rel_move_subst B δ σ v w : + 𝒱 B (λ n, interp_type (σ n) δ) v w ↔ 𝒱 (B.[σ]) δ v w. + Proof. + induction B in σ, δ, v, w |-*; simpl; simp type_interp; eauto. + - f_equiv; intros e1. f_equiv; intros e2. do 4 f_equiv. + eapply forall_proper; intros τ. + simp type_interp. f_equiv. intros v1. f_equiv. intros v2. + do 2 f_equiv. etransitivity; last apply IHB. + eapply sem_val_rel_ext; intros [|n] u. + + simp type_interp. done. + + simpl. rewrite /up. simpl. + etransitivity; last eapply sem_val_rel_move_ren. + simpl. done. + - f_equiv. intros v1. f_equiv. intros v2. do 3 f_equiv. intros τ. + etransitivity; last apply IHB. + eapply sem_val_rel_ext; intros [|n] u. + + simp type_interp. done. + + simpl. rewrite /up. simpl. + etransitivity; last eapply sem_val_rel_move_ren. + simpl. done. + - do 4 (f_equiv; intros ?). + do 4 f_equiv. + eapply forall_proper; intros v1. + eapply forall_proper; intros v2. + eapply if_iff; first eauto. + simp type_interp. f_equiv. intros v3. + f_equiv. intros v4. + do 2 f_equiv. + eauto. + - do 4 (f_equiv; intros ?). + do 3 f_equiv; eauto. + - f_equiv; f_equiv; intros ?; f_equiv; intros ?. + all: do 2 f_equiv; eauto. + Qed. + + Lemma sem_val_rel_move_single_subst A B δ v w : + 𝒱 B (interp_type A δ .: δ) v w ↔ 𝒱 (B.[A/]) δ v w. + Proof. + rewrite -sem_val_rel_move_subst. eapply sem_val_rel_ext. + intros [| n] w1 w2; simpl; done. + Qed. + + Lemma sem_val_rel_cons A δ v w τ : + 𝒱 A δ v w ↔ 𝒱 A.[ren (+1)] (τ .: δ) v w. + Proof. + rewrite -sem_val_rel_move_subst; simpl. + eapply sem_val_rel_ext; done. + Qed. + + Lemma sem_context_rel_cons Γ δ θ1 θ2 τ : + 𝒢 δ Γ θ1 θ2 → + 𝒢 (τ .: δ) (⤉ Γ) θ1 θ2. + Proof. + induction 1 as [ | Γ θ1 θ2 v w x A Hv Hctx IH]; simpl. + - rewrite fmap_empty. constructor. + - rewrite fmap_insert. constructor; last done. + rewrite -sem_val_rel_cons. done. + Qed. +End boring_lemmas. + +(** ** Compatibility lemmas *) + +Lemma compat_int Δ Γ z : TY Δ; Γ ⊨ (Lit $ LitInt z) ≈ (Lit $ LitInt z) : Int. +Proof. + do 2 (split; first done). + intros θ1 θ2 δ _. simp type_interp. + exists #z, #z. + split; first by constructor. + split; first by constructor. + simp type_interp. eauto. +Qed. + +Lemma compat_bool Δ Γ b : TY Δ; Γ ⊨ (Lit $ LitBool b) ≈ (Lit $ LitBool b) : Bool. +Proof. + do 2 (split; first done). + intros θ1 θ2 δ _. simp type_interp. + exists #b, #b. + split; first by constructor. + split; first by constructor. + simp type_interp. eauto. +Qed. + +Lemma compat_unit Δ Γ : TY Δ; Γ ⊨ (Lit $ LitUnit) ≈ (Lit $ LitUnit) : Unit. +Proof. + do 2 (split; first done). + intros θ1 θ2 δ _. simp type_interp. + exists #LitUnit, #LitUnit. + split; first by constructor. + split; first by constructor. + simp type_interp. split; eauto. +Qed. + +Lemma compat_var Δ Γ x A : + Γ !! x = Some A → + TY Δ; Γ ⊨ (Var x) ≈ (Var x) : A. +Proof. + intros Hx. + do 2 (split; first eapply bool_decide_pack, elem_of_elements, elem_of_dom_2, Hx). + intros θ1 θ2 δ Hctx; simpl. + + specialize (sem_context_rel_vals Hctx Hx) as (e1 & e2 & v1 & v2 & He1 & He2 & Heq1 & Heq2 & Hv). + rewrite He1 He2. simp type_interp. exists v1, v2. + repeat split; last done. + - rewrite -(of_to_val _ _ Heq1). + by apply big_step_of_val. + - rewrite -(of_to_val _ _ Heq2). + by apply big_step_of_val. +Qed. + +Lemma compat_app Δ Γ e1 e1' e2 e2' A B : + TY Δ; Γ ⊨ e1 ≈ e1': (A → B) → + TY Δ; Γ ⊨ e2 ≈ e2' : A → + TY Δ; Γ ⊨ (e1 e2) ≈ (e1' e2') : B. +Proof. + intros (Hfuncl & Hfuncl' & Hfun) (Hargcl & Hargcl' & Harg). + do 2 (split; first naive_solver). + intros θ1 θ2 δ Hctx; simpl. + + specialize (Hfun _ _ _ Hctx). simp type_interp in Hfun. destruct Hfun as (v1 & v2 & Hbs1 & Hbs2 & Hv12). + simp type_interp in Hv12. destruct Hv12 as (x & y & e1'' & e2'' & -> & -> & ? & ? & Hv12). + specialize (Harg _ _ _ Hctx). simp type_interp in Harg. + destruct Harg as (v3 & v4 & Hbs3 & Hbs4 & Hv34). + + apply Hv12 in Hv34. + simp type_interp in Hv34. + destruct Hv34 as (v5 & v6 & Hbs5 & Hbs6 & Hv56). + + simp type_interp. + exists v5, v6. eauto. +Qed. + + + +(* Compatibility for [lam] unfortunately needs a very technical helper lemma. *) +Lemma lam_closed Γ θ (x : string) A e : + dom Γ = dom θ → + subst_is_closed [] θ → + closed (elements (dom (<[x:=A]> Γ))) e → + closed [] (Lam x (subst_map (delete x θ) e)). +Proof. + intros Hdom Hsubstcl Hcl. + eapply subst_map_closed. + - eapply is_closed_weaken; first done. + rewrite dom_delete dom_insert Hdom //. + intros y. destruct (decide (x = y)); set_solver. + - intros x' e' Hx. + eapply (is_closed_weaken []); last set_solver. + eapply Hsubstcl. + eapply map_subseteq_spec; last done. + apply map_delete_subseteq. +Qed. +(** Lambdas need to be closed by the context *) +Lemma compat_lam_named Δ Γ x e1 e2 A B : + TY Δ; (<[ x := A ]> Γ) ⊨ e1 ≈ e2 : B → + TY Δ; Γ ⊨ (Lam (BNamed x) e1) ≈ (Lam (BNamed x) e2): (A → B). +Proof. + intros (Hbodycl & Hbodycl' & Hbody). + do 2 (split; first (simpl; eapply is_closed_weaken; set_solver)). + intros θ1 θ2 δ Hctx. simpl. + simp type_interp. + + exists ((λ: x, subst_map (delete x θ1) e1))%V, ((λ: x, subst_map (delete x θ2) e2))%V. + split; first by eauto. + split; first by eauto. + simp type_interp. + + opose proof* sem_context_rel_dom as []; first done. + opose proof* sem_context_rel_closed as []; first done. + eexists (BNamed x), (BNamed x), _, _. split_and!. + 1-2: reflexivity. + 1-2: eapply lam_closed; eauto. + intros v' w' Hvw'. + specialize (Hbody (<[ x := of_val v']> θ1) (<[ x := of_val w']> θ2)). + simpl. generalize Hctx=>Hctx'. + eapply sem_context_rel_closed in Hctx' as Hclosed. + rewrite !subst_subst_map; [|naive_solver..]. + apply Hbody. apply sem_context_rel_insert; done. +Qed. + + +Lemma compat_lam_anon Δ Γ e1 e2 A B : + TY Δ; Γ ⊨ e1 ≈ e2 : B → + TY Δ; Γ ⊨ (Lam BAnon e1) ≈ (Lam BAnon e2) : (A → B). +Proof. + intros (Hbodycl & Hbodycl' & Hbody). + do 2 (split; first done). + intros θ1 θ2 δ Hctx. simpl. + simp type_interp. + + exists (λ: <>, subst_map θ1 e1)%V, (λ: <>, subst_map θ2 e2)%V. + split; first by eauto. + split; first by eauto. + simp type_interp. + eexists BAnon, BAnon, _, _. split_and!; try reflexivity. + - simpl. eapply subst_map_closed; simpl. + + replace (dom θ1) with (dom Γ); first done. + eapply sem_context_rel_dom in Hctx. naive_solver. + + apply sem_context_rel_closed in Hctx. naive_solver. + - simpl. eapply subst_map_closed; simpl. + + replace (dom θ2) with (dom Γ); first done. + eapply sem_context_rel_dom in Hctx. naive_solver. + + apply sem_context_rel_closed in Hctx. naive_solver. + - intros v' w' Hvw'. + specialize (Hbody θ1 θ2). + simpl. apply Hbody; done. +Qed. + +Lemma compat_int_binop Δ Γ op e1 e1' e2 e2' : + bin_op_typed op Int Int Int → + TY Δ; Γ ⊨ e1 ≈ e1' : Int → + TY Δ; Γ ⊨ e2 ≈ e2' : Int → + TY Δ; Γ ⊨ (BinOp op e1 e2) ≈ (BinOp op e1' e2') : Int. +Proof. + intros Hop (He1cl & He1cl' & He1) (He2cl & He2cl' & He2). + do 2 (split; first naive_solver). + intros θ1 θ2 δ Hctx. simpl. + simp type_interp. + specialize (He1 _ _ _ Hctx). specialize (He2 _ _ _ Hctx). + simp type_interp in He1. simp type_interp in He2. + + destruct He1 as (v1 & v1' & Hb1 & Hb1' & Hv1). + destruct He2 as (v2 & v2' & Hb2 & Hb2' & Hv2). + simp type_interp in Hv1, Hv2. + destruct Hv1 as (z1 & -> & ->). + destruct Hv2 as (z2 & -> & ->). + + inversion Hop; subst op. + + exists #(z1 + z2)%Z, #(z1 + z2)%Z. split_and!. + - econstructor; done. + - econstructor; done. + - exists (z1 + z2)%Z. done. + + exists #(z1 - z2)%Z, #(z1 - z2)%Z. split_and!. + - econstructor; done. + - econstructor; done. + - exists (z1 - z2)%Z. done. + + exists #(z1 * z2)%Z, #(z1 * z2)%Z. split_and!. + - econstructor; done. + - econstructor; done. + - exists (z1 * z2)%Z. done. +Qed. + +Lemma compat_int_bool_binop Δ Γ op e1 e1' e2 e2' : + bin_op_typed op Int Int Bool → + TY Δ; Γ ⊨ e1 ≈ e1' : Int → + TY Δ; Γ ⊨ e2 ≈ e2' : Int → + TY Δ; Γ ⊨ (BinOp op e1 e2) ≈ (BinOp op e1' e2') : Bool. +Proof. + intros Hop (He1cl & He1cl' & He1) (He2cl & He2cl' & He2). + do 2 (split; first naive_solver). + intros θ1 θ2 δ Hctx. simpl. + + simp type_interp. + specialize (He1 _ _ _ Hctx). specialize (He2 _ _ _ Hctx). + simp type_interp in He1. simp type_interp in He2. + + destruct He1 as (v1 & v1' & Hb1 & Hb1' & Hv1). + destruct He2 as (v2 & v2' & Hb2 & Hb2' & Hv2). + simp type_interp in Hv1, Hv2. + destruct Hv1 as (z1 & -> & ->). + destruct Hv2 as (z2 & -> & ->). + + inversion Hop; subst op. + + exists #(bool_decide (z1 < z2))%Z, #(bool_decide (z1 < z2))%Z. split_and!. + - econstructor; done. + - econstructor; done. + - by eexists _. + + exists #(bool_decide (z1 ≤ z2))%Z, #(bool_decide (z1 ≤ z2))%Z. split_and!. + - econstructor; done. + - econstructor; done. + - by eexists _. + + exists #(bool_decide (z1 = z2))%Z, #(bool_decide (z1 = z2))%Z. split_and!. + - econstructor; done. + - econstructor; done. + - eexists _. done. +Qed. + +Lemma compat_unop Δ Γ op A B e e' : + un_op_typed op A B → + TY Δ; Γ ⊨ e ≈ e' : A → + TY Δ; Γ ⊨ (UnOp op e) ≈ (UnOp op e') : B. +Proof. + intros Hop (Hecl & Hecl' & He). + do 2 (split; first naive_solver). + intros θ1 θ2 δ Hctx. simpl. + + simp type_interp. specialize (He _ _ _ Hctx). + simp type_interp in He. + + destruct He as (v & v' & Hb & Hb' & Hv). + inversion Hop; subst; simp type_interp in Hv. + - destruct Hv as (b & -> & ->). + exists #(negb b), #(negb b). split_and!. + + econstructor; done. + + econstructor; done. + + by eexists _. + - destruct Hv as (z & -> & ->). + exists #(-z)%Z, #(-z)%Z. split_and!. + + econstructor; done. + + econstructor; done. + + by eexists _. +Qed. + + +Lemma compat_tlam Δ Γ e1 e2 A : + TY S Δ; (⤉ Γ) ⊨ e1 ≈ e2 : A → + TY Δ; Γ ⊨ (Λ, e1) ≈ (Λ, e2) : (∀: A). +Proof. + intros (Hcl & Hcl' & He). + do 2 (split; first (simpl; by erewrite <-dom_fmap)). + + intros θ1 θ2 δ Hctx. simpl. + simp type_interp. + exists (Λ, subst_map θ1 e1)%V, (Λ, subst_map θ2 e2)%V. + split; first constructor. + split; first constructor. + + simp type_interp. + eexists _, _. split_and!; try done. + - simpl. eapply subst_map_closed; simpl. + + replace (dom θ1) with (dom Γ). + * by erewrite <-dom_fmap. + * apply sem_context_rel_dom in Hctx. naive_solver. + + apply sem_context_rel_closed in Hctx. naive_solver. + - simpl. eapply subst_map_closed; simpl. + + replace (dom θ2) with (dom Γ). + * by erewrite <-dom_fmap. + * apply sem_context_rel_dom in Hctx. naive_solver. + + apply sem_context_rel_closed in Hctx. naive_solver. + - intros τ. eapply He. + by eapply sem_context_rel_cons. +Qed. + +Lemma compat_tapp Δ Γ e e' A B : + type_wf Δ B → + TY Δ; Γ ⊨ e ≈ e' : (∀: A) → + TY Δ; Γ ⊨ (e <>) ≈ (e' <>) : (A.[B/]). +Proof. + intros Hwf (Hecl & Hecl' & He). + do 2 (split; first naive_solver). + intros θ1 θ2 δ Hctx. simpl. + + simp type_interp. + specialize (He _ _ _ Hctx). simp type_interp in He. + destruct He as (v & v' & Hb & Hb' & Hv). + simp type_interp in Hv. + destruct Hv as (e1 & e2 & -> & -> & Cl & Cl' & He1). + + set (τ := interp_type B δ). + specialize (He1 τ). + simp type_interp in He1. destruct He1 as (v & w & Hb2 & Hb2' & Hv). + exists v, w. split_and!; try by econstructor. + apply sem_val_rel_move_single_subst. done. +Qed. + +Lemma compat_pack Δ Γ e e' n A B : + type_wf n B → + type_wf (S n) A → + TY n; Γ ⊨ e ≈ e': A.[B/] → + TY n; Γ ⊨ (pack e) ≈ (pack e') : (∃: A). +Proof. + intros Hwf Hwf' (Hecl & Hecl' & He). + do 2 (split; first naive_solver). + intros θ1 θ2 δ Hctx. simpl. + + simp type_interp. + + specialize (He _ _ _ Hctx). simp type_interp in He. + destruct He as (v & v' & Hb & Hb' & Hv). + exists (PackV v), (PackV v'). + split; first eauto. + split; first eauto. + + simp type_interp. exists v, v'. + split; first done. + split; first done. + exists (interp_type B δ). + apply sem_val_rel_move_single_subst. done. +Qed. + +Lemma compat_unpack n Γ A B e1 e1' e2 e2' x : + type_wf n B → + TY n; Γ ⊨ e1 ≈ e2 : (∃: A) → + TY S n; <[x:=A]> (⤉Γ) ⊨ e1' ≈ e2' : B.[ren (+1)] → + TY n; Γ ⊨ (unpack e1 as BNamed x in e1') ≈ (unpack e2 as BNamed x in e2') : B. +Proof. + intros Hwf (Hecl & Hecl' & He) (He'cl & He'cl' & He'). + split. { simpl. apply andb_True. split; first done. + eapply is_closed_weaken; set_solver. } + split. { simpl. apply andb_True. split; first done. + eapply is_closed_weaken; set_solver. } + intros θ1 θ2 δ Hctx. simpl. + + simp type_interp. + specialize (He _ _ _ Hctx). simp type_interp in He. + destruct He as (v & w & Hb & Hb' & Hv). + simp type_interp in Hv. destruct Hv as (v' & w' & -> & -> & τ & Hv'). + + specialize (He' (<[x := of_val v']> θ1) (<[x := of_val w']> θ2) (τ.:δ)). + simp type_interp in He'. + destruct He' as (v & w & Hb'' & Hb''' & Hv). + { constructor; first done. by apply sem_context_rel_cons. } + exists v, w. split_and!. + - econstructor; first done. rewrite subst'_subst_map; first done. + eapply sem_context_rel_closed in Hctx. naive_solver. + - econstructor; first done. rewrite subst'_subst_map; first done. + eapply sem_context_rel_closed in Hctx. naive_solver. + - by eapply sem_val_rel_cons. +Qed. + +Lemma compat_if n Γ e0 e0' e1 e1' e2 e2' A : + TY n; Γ ⊨ e0 ≈ e0' : Bool → + TY n; Γ ⊨ e1 ≈ e1' : A → + TY n; Γ ⊨ e2 ≈ e2' : A → + TY n; Γ ⊨ (if: e0 then e1 else e2) ≈ (if: e0' then e1' else e2') : A. +Proof. + intros (He0cl & He0cl' & He0) (He1cl & He1cl' & He1) (He2cl & He2cl' & He2). + do 2 (split; first naive_solver). + intros θ1 θ2 δ Hctx. simpl. + simp type_interp. + + specialize (He0 _ _ _ Hctx). simp type_interp in He0. + specialize (He1 _ _ _ Hctx). simp type_interp in He1. + specialize (He2 _ _ _ Hctx). simp type_interp in He2. + + destruct He0 as (v0 & v0' & Hb0 & Hb0' & Hv0). simp type_interp in Hv0. + destruct Hv0 as (b & -> & ->). + destruct He1 as (v1 & w1 & Hb1 & Hb1' & Hv1). + destruct He2 as (v2 & w2 & Hb2 & Hb2' & Hv2). + + destruct b. + - exists v1, w1. split_and!; try by repeat econstructor. + - exists v2, w2. split_and!; try by repeat econstructor. +Qed. + +Lemma compat_pair Δ Γ e1 e1' e2 e2' A B : + TY Δ; Γ ⊨ e1 ≈ e1' : A → + TY Δ; Γ ⊨ e2 ≈ e2' : B → + TY Δ; Γ ⊨ (e1, e2) ≈ (e1', e2') : A × B. +Proof. + intros (He1cl & He1cl' & He1) (He2cl & He2cl' & He2). + do 2 (split; first naive_solver). + intros θ1 θ2 δ Hctx. simpl. + simp type_interp. + + specialize (He1 _ _ _ Hctx). specialize (He2 _ _ _ Hctx). + simp type_interp in He1. simp type_interp in He2. + + destruct He1 as (v1 & v1' & Hb1 & Hb1' & Hv1). + destruct He2 as (v2 & v2' & Hb2 & Hb2' & Hv2). + simp type_interp in Hv1, Hv2. + eexists _, _. split_and!; try by econstructor. + simp type_interp. eexists _, _, _, _. + split_and!; eauto. +Qed. + +Lemma compat_fst Δ Γ e e' A B : + TY Δ; Γ ⊨ e ≈ e' : A × B → + TY Δ; Γ ⊨ Fst e ≈ Fst e' : A. +Proof. + intros (Hecl & Hecl' & He). + do 2 (split; first naive_solver). + intros θ1 θ2 δ Hctx. simpl. + simp type_interp. + + specialize (He _ _ _ Hctx). simp type_interp in He. + destruct He as (v & v' & Hb & Hb' & Hv). + simp type_interp in Hv. + destruct Hv as (v1 & v2 & w1 & w2 & -> & -> & Hv & Hw). + eexists _, _. + split_and!; eauto. +Qed. + +Lemma compat_snd Δ Γ e e' A B : + TY Δ; Γ ⊨ e ≈ e' : A × B → + TY Δ; Γ ⊨ Snd e ≈ Snd e' : B. +Proof. + intros (Hecl & Hecl' & He). + do 2 (split; first naive_solver). + intros θ1 θ2 δ Hctx. simpl. + simp type_interp. + + specialize (He _ _ _ Hctx). simp type_interp in He. + destruct He as (v & v' & Hb & Hb' & Hv). + simp type_interp in Hv. + destruct Hv as (v1 & v2 & w1 & w2 & -> & -> & Hv & Hw). + eexists _, _. + split_and!; eauto. +Qed. + +Lemma compat_injl Δ Γ e e' A B : + TY Δ; Γ ⊨ e ≈ e' : A → + TY Δ; Γ ⊨ InjL e ≈ InjL e' : A + B. +Proof. + intros (Hecl & Hecl' & He). + do 2 (split; first naive_solver). + intros θ1 θ2 δ Hctx. simpl. + simp type_interp. + + specialize (He _ _ _ Hctx). simp type_interp in He. + destruct He as (v & v' & Hb & Hb' & Hv). + simp type_interp in Hv. + eexists _, _. + split_and!; eauto. + simp type_interp. + left. eexists _, _. split_and!; eauto. +Qed. + +Lemma compat_injr Δ Γ e e' A B : + TY Δ; Γ ⊨ e ≈ e' : B → + TY Δ; Γ ⊨ InjR e ≈ InjR e' : A + B. +Proof. + intros (Hecl & Hecl' & He). + do 2 (split; first naive_solver). + intros θ1 θ2 δ Hctx. simpl. + simp type_interp. + + specialize (He _ _ _ Hctx). simp type_interp in He. + destruct He as (v & v' & Hb & Hb' & Hv). + simp type_interp in Hv. + eexists _, _. + split_and!; eauto. + simp type_interp. + right. eexists _, _. split_and!; eauto. +Qed. + +Lemma compat_case Δ Γ e e' e1 e1' e2 e2' A B C : + TY Δ; Γ ⊨ e ≈ e' : B + C → + TY Δ; Γ ⊨ e1 ≈ e1' : (B → A) → + TY Δ; Γ ⊨ e2 ≈ e2' : (C → A) → + TY Δ; Γ ⊨ Case e e1 e2 ≈ Case e' e1' e2' : A. +Proof. + intros (He0cl & He0cl' & He0) (He1cl & He1cl' & He1) (He2cl & He2cl' & He2). + do 2 (split; first naive_solver). + intros θ1 θ2 δ Hctx. simpl. + simp type_interp. + + specialize (He0 _ _ _ Hctx). simp type_interp in He0. + specialize (He1 _ _ _ Hctx). simp type_interp in He1. + specialize (He2 _ _ _ Hctx). simp type_interp in He2. + + destruct He0 as (v0 & v0' & Hb0 & Hb0' & Hv0). simp type_interp in Hv0. + destruct He1 as (v1 & w1 & Hb1 & Hb1' & Hv1). + destruct He2 as (v2 & w2 & Hb2 & Hb2' & Hv2). + + destruct Hv0 as [(v' & w' & -> & -> & Hv)|(v' & w' & -> & -> & Hv)]. + - simp type_interp in Hv1. destruct Hv1 as (x & y & e'' & e''' & -> & -> & Cl & Cl' & Hv1). + apply Hv1 in Hv. simp type_interp in Hv. destruct Hv as (v & w & Hb''' & Hb'''' & Hv''). + eexists _, _. split_and!; eauto using big_step, big_step_of_val. + - simp type_interp in Hv2. destruct Hv2 as (x & y & e'' & e''' & -> & -> & Cl & Cl' & Hv2). + apply Hv2 in Hv. simp type_interp in Hv. destruct Hv as (v & w & Hb''' & Hb'''' & Hv''). + eexists _, _. split_and!; eauto using big_step, big_step_of_val. +Qed. + + + +(* we register the compatibility lemmas with eauto *) +Local Hint Resolve + compat_var compat_lam_named compat_lam_anon + compat_tlam compat_tapp compat_pack compat_unpack + compat_int compat_bool compat_unit compat_if + compat_app compat_int_binop compat_int_bool_binop + compat_unop compat_pair compat_fst compat_snd + compat_injl compat_injr compat_case : core. + + + + + +Lemma sem_soundness Δ Γ e A : + TY Δ; Γ ⊢ e : A → + TY Δ; Γ ⊨ e ≈ e : A. +Proof. + induction 1 as [ | Δ Γ x e A B Hsyn IH | Δ Γ e A B Hsyn IH| Δ Γ e A Hsyn IH| | | | | | | | | n Γ e1 e2 op A B C [] ? ? ? ? | | | | | | | ]; eauto. +Qed. + + +Program Definition any_type : sem_type := {| sem_type_car := λ v w, is_closed [] v ∧ is_closed [] w |}. +Definition δ_any : var → sem_type := λ _, any_type. + + + +(* Contextual Equivalence *) +Inductive pctx := + | HolePCtx + | AppLPCtx (C: pctx) (e2 : expr) + | AppRPCtx (e1 : expr) (C: pctx) + | TAppPCtx (C: pctx) + | PackPCtx (C: pctx) + | UnpackLPCtx (x : binder)(C: pctx) (e2 : expr) + | UnpackRPCtx (x : binder) (e1 : expr) (C: pctx) + | UnOpPCtx (op : un_op) (C: pctx) + | BinOpLPCtx (op : bin_op) (C: pctx) (e2 : expr) + | BinOpRPCtx (op : bin_op) (e1 : expr) (C: pctx) + | IfPCtx (C: pctx) (e1 e2 : expr) + | IfTPCtx (e: expr) (C: pctx) (e2 : expr) + | IfEPCtx (e e1: expr) (C: pctx) + | PairLPCtx (C: pctx) (e2 : expr) + | PairRPCtx (e1 : expr) (C: pctx) + | FstPCtx (C: pctx) + | SndPCtx (C: pctx) + | InjLPCtx (C: pctx) + | InjRPCtx (C: pctx) + | CasePCtx (C: pctx) (e1 e2 : expr) + | CaseTPCtx (e: expr) (C: pctx) (e2 : expr) + | CaseEPCtx (e e1: expr) (C: pctx) + | LamPCtx (x: binder) (C: pctx) + | TLamPCtx (C: pctx). + +Fixpoint pfill (C : pctx) (e : expr) : expr := + match C with + | HolePCtx => e + | AppLPCtx K e2 => App (pfill K e) e2 + | AppRPCtx e1 K => App e1 (pfill K e) + | TAppPCtx K => TApp (pfill K e) + | PackPCtx K => Pack (pfill K e) + | UnpackLPCtx x K e2 => Unpack x (pfill K e) e2 + | UnpackRPCtx x e1 K => Unpack x e1 (pfill K e) + | UnOpPCtx op K => UnOp op (pfill K e) + | BinOpLPCtx op K e2 => BinOp op (pfill K e) e2 + | BinOpRPCtx op e1 K => BinOp op e1 (pfill K e) + | IfPCtx K e1 e2 => If (pfill K e) e1 e2 + | IfTPCtx e' K e2 => If e' (pfill K e) e2 + | IfEPCtx e' e1 K => If e' e1 (pfill K e) + | PairLPCtx K e2 => Pair (pfill K e) e2 + | PairRPCtx e1 K => Pair e1 (pfill K e) + | FstPCtx K => Fst (pfill K e) + | SndPCtx K => Snd (pfill K e) + | InjLPCtx K => InjL (pfill K e) + | InjRPCtx K => InjR (pfill K e) + | CasePCtx K e1 e2 => Case (pfill K e) e1 e2 + | CaseTPCtx e' K e2 => Case e' (pfill K e) e2 + | CaseEPCtx e' e1 K => Case e' e1 (pfill K e) + | LamPCtx x C => Lam x (pfill C e) + | TLamPCtx C => TLam (pfill C e) + end. + + +Inductive pctx_typed (Δ: nat) (Γ: typing_context) (A: type): pctx → nat → typing_context → type → Prop := + | pctx_typed_HolePCtx : pctx_typed Δ Γ A HolePCtx Δ Γ A + | pctx_typed_AppLPCtx K e2 B C Δ' Γ': + pctx_typed Δ Γ A K Δ' Γ' (B → C) → + TY Δ'; Γ' ⊢ e2 : B → + pctx_typed Δ Γ A (AppLPCtx K e2) Δ' Γ' C + | pctx_typed_AppRPCtx e1 K B C Δ' Γ': + (TY Δ'; Γ' ⊢ e1 : Fun B C) → + pctx_typed Δ Γ A K Δ' Γ' B → + pctx_typed Δ Γ A (AppRPCtx e1 K) Δ' Γ' C + | pctx_typed_TLamPCtx K B Δ' Γ': + pctx_typed Δ Γ A K (S Δ') (⤉ Γ') B → + pctx_typed Δ Γ A (TLamPCtx K) Δ' Γ' (∀: B) + | pctx_typed_TAppPCtx K B C Δ' Γ': + type_wf Δ' C → + pctx_typed Δ Γ A K Δ' Γ' (∀: B) → + pctx_typed Δ Γ A (TAppPCtx K) Δ' Γ' (B.[C/]) + | pctx_typed_PackPCtx K B C Δ' Γ': + type_wf Δ' C → + type_wf (S Δ') B → + pctx_typed Δ Γ A K Δ' Γ' (B.[C/]) → + pctx_typed Δ Γ A (PackPCtx K) Δ' Γ' (∃: B) + | pctx_typed_UnpackLPCtx (x: string) K e2 B C Δ' Γ' : + type_wf Δ' C → + pctx_typed Δ Γ A K Δ' Γ' (∃: B) → + (TY S Δ'; (<[x := B]> (⤉ Γ')) ⊢ e2 : C.[ren (+1)]) → + pctx_typed Δ Γ A (UnpackLPCtx x K e2) Δ' Γ' C + | pctx_typed_UnpackRPCtx (x: string) e1 K B C Δ' Γ' : + type_wf Δ' C → + (TY Δ'; Γ' ⊢ e1 : (∃: B)) → + (pctx_typed Δ Γ A K (S Δ') (<[x := B]> (⤉ Γ')) (C.[ren (+1)])) → + pctx_typed Δ Γ A (UnpackRPCtx x e1 K) Δ' Γ' C + | pctx_typed_UnOpPCtx op K Δ' Γ' B C: + un_op_typed op B C → + pctx_typed Δ Γ A K Δ' Γ' B → + pctx_typed Δ Γ A (UnOpPCtx op K) Δ' Γ' C + | pctx_typed_BinOpLPCtx op K e2 B1 B2 C Δ' Γ' : + bin_op_typed op B1 B2 C → + pctx_typed Δ Γ A K Δ' Γ' B1 → + TY Δ'; Γ' ⊢ e2 : B2 → + pctx_typed Δ Γ A (BinOpLPCtx op K e2) Δ' Γ' C + | pctx_typed_BinOpRPCtx op K e1 B1 B2 C Δ' Γ' : + bin_op_typed op B1 B2 C → + TY Δ'; Γ' ⊢ e1 : B1 → + pctx_typed Δ Γ A K Δ' Γ' B2 → + pctx_typed Δ Γ A (BinOpRPCtx op e1 K) Δ' Γ' C + | pctx_typed_IfPCtx K e1 e2 B Δ' Γ' : + pctx_typed Δ Γ A K Δ' Γ' Bool → + TY Δ'; Γ' ⊢ e1 : B → + TY Δ'; Γ' ⊢ e2 : B → + pctx_typed Δ Γ A (IfPCtx K e1 e2) Δ' Γ' B + | pctx_typed_IfTPCtx K e e2 B Δ' Γ' : + TY Δ'; Γ' ⊢ e : Bool → + pctx_typed Δ Γ A K Δ' Γ' B → + TY Δ'; Γ' ⊢ e2 : B → + pctx_typed Δ Γ A (IfTPCtx e K e2) Δ' Γ' B + | pctx_typed_IfEPCtx K e e1 B Δ' Γ' : + TY Δ'; Γ' ⊢ e : Bool → + TY Δ'; Γ' ⊢ e1 : B → + pctx_typed Δ Γ A K Δ' Γ' B → + pctx_typed Δ Γ A (IfEPCtx e e1 K) Δ' Γ' B + | pctx_typed_PairLPCtx K e2 B1 B2 Δ' Γ' : + pctx_typed Δ Γ A K Δ' Γ' B1 → + TY Δ'; Γ' ⊢ e2 : B2 → + pctx_typed Δ Γ A (PairLPCtx K e2) Δ' Γ' (Prod B1 B2) + | pctx_typed_PairRPCtx K e1 B1 B2 Δ' Γ' : + TY Δ'; Γ' ⊢ e1 : B1 → + pctx_typed Δ Γ A K Δ' Γ' B2 → + pctx_typed Δ Γ A (PairRPCtx e1 K) Δ' Γ' (Prod B1 B2) + | pctx_typed_FstPCtx K Δ' Γ' B C: + pctx_typed Δ Γ A K Δ' Γ' (Prod B C) → + pctx_typed Δ Γ A (FstPCtx K) Δ' Γ' B + | pctx_typed_SndPCtx K Δ' Γ' B C: + pctx_typed Δ Γ A K Δ' Γ' (Prod B C) → + pctx_typed Δ Γ A (SndPCtx K) Δ' Γ' C + | pctx_typed_InjLPCtx K Δ' Γ' B C: + type_wf Δ' C → + pctx_typed Δ Γ A K Δ' Γ' B → + pctx_typed Δ Γ A (InjLPCtx K) Δ' Γ' (Sum B C) + | pctx_typed_InjRPCtx K Δ' Γ' B C: + type_wf Δ' B → + pctx_typed Δ Γ A K Δ' Γ' C → + pctx_typed Δ Γ A (InjRPCtx K) Δ' Γ' (Sum B C) + | pctx_typed_CasePCtx K e1 e2 B C D Δ' Γ' : + pctx_typed Δ Γ A K Δ' Γ' (Sum B C) → + TY Δ'; Γ' ⊢ e1 : (Fun B D) → + TY Δ'; Γ' ⊢ e2 : (Fun C D) → + pctx_typed Δ Γ A (CasePCtx K e1 e2) Δ' Γ' D + | pctx_typed_CaseTPCtx K e e2 B C D Δ' Γ' : + TY Δ'; Γ' ⊢ e : (Sum B C) → + pctx_typed Δ Γ A K Δ' Γ' (Fun B D) → + TY Δ'; Γ' ⊢ e2 : (Fun C D) → + pctx_typed Δ Γ A (CaseTPCtx e K e2) Δ' Γ' D + | pctx_typed_CaseEPCtx K e e1 B C D Δ' Γ' : + TY Δ'; Γ' ⊢ e : (Sum B C) → + TY Δ'; Γ' ⊢ e1 : (Fun B D) → + pctx_typed Δ Γ A K Δ' Γ' (Fun C D) → + pctx_typed Δ Γ A (CaseEPCtx e e1 K) Δ' Γ' D + | pctx_typed_named_LamPCtx (x: string) K B C Γ' Δ' : + type_wf Δ' B → + pctx_typed Δ Γ A K Δ' (<[x := B]> Γ') C → + pctx_typed Δ Γ A (LamPCtx x K) Δ' Γ' (Fun B C) + | pctx_typed_anon_LamPCtx K B C Γ' Δ' : + type_wf Δ' B → + pctx_typed Δ Γ A K Δ' Γ' C → + pctx_typed Δ Γ A (LamPCtx BAnon K) Δ' Γ' (Fun B C) + . + + +Lemma pfill_typed C Δ Δ' Γ Γ' e A B: + pctx_typed Δ Γ A C Δ' Γ' B → TY Δ; Γ ⊢ e : A → TY Δ'; Γ' ⊢ pfill C e : B. +Proof. + induction 1 in |-*; simpl; eauto using pctx_typed. +Qed. + + +Lemma syn_typed_closed Δ Γ A e: + TY Δ; Γ ⊢ e : A → + is_closed (elements (dom Γ)) e. +Proof. + intros Hty; eapply syn_typed_closed; eauto. + intros x Hx. by eapply elem_of_elements. +Qed. + +Lemma pctx_typed_fill_closed Δ Δ' Γ Γ' A B K e: + is_closed (elements (dom Γ)) e → + pctx_typed Δ Γ A K Δ' Γ' B → + is_closed (elements (dom Γ')) (pfill K e). +Proof. + intros Hcl. induction 1; simplify_closed; eauto using syn_typed_closed. + - eapply is_closed_weaken; first by eapply syn_typed_closed. + rewrite dom_insert. + intros y Hin. destruct (decide (x = y)); subst; first set_solver. + eapply elem_of_elements in Hin. eapply elem_of_union in Hin as []. + + set_solver. + + rewrite dom_fmap in H2. eapply elem_of_list_further. + by eapply elem_of_elements. + - rewrite dom_insert. + intros y Hin. destruct (decide (x = y)); subst; first set_solver. + eapply elem_of_elements in Hin. eapply elem_of_union in Hin as []. + + set_solver. + + rewrite dom_fmap in H2. eapply elem_of_list_further. + by eapply elem_of_elements. + - rewrite dom_insert. + intros y Hin. destruct (decide (x = y)); subst; first set_solver. + eapply elem_of_elements in Hin. eapply elem_of_union in Hin as []. + + set_solver. + + eapply elem_of_list_further. + by eapply elem_of_elements. +Qed. + + +Lemma sem_typed_congruence Δ Δ' Γ Γ' e1 e2 C A B : + closed (elements (dom Γ)) e1 → + closed (elements (dom Γ)) e2 → + TY Δ; Γ ⊨ e1 ≈ e2 : A → + pctx_typed Δ Γ A C Δ' Γ' B → + TY Δ'; Γ' ⊨ pfill C e1 ≈ pfill C e2 : B. +Proof. + intros ???. + induction 1; simpl; eauto using sem_soundness. + - inversion H2; subst; eauto using sem_soundness. + - inversion H2; subst; eauto using sem_soundness. +Qed. + +Lemma adequacy δ e1 e2: ℰ Int δ e1 e2 → ∃ n, big_step e1 n ∧ big_step e2 n. +Proof. + simp type_interp. intros (? & ? & ? & ? & Hty). + simp type_interp in Hty. naive_solver. +Qed. + + +Definition ctx_equiv Δ Γ e1 e2 A := + ∀ K, pctx_typed Δ Γ A K 0 ∅ Int → ∃ n: Z, big_step (pfill K e1) #n ∧ big_step (pfill K e2) #n. + + +Lemma sem_typing_ctx_equiv Δ Γ e1 e2 A : + (* the closedness assumptions could be replaced by typing assumptions *) + closed (elements (dom Γ)) e1 → + closed (elements (dom Γ)) e2 → + TY Δ; Γ ⊨ e1 ≈ e2 : A → ctx_equiv Δ Γ e1 e2 A. +Proof. + intros Hcl Hcl' Hsem C Hty. eapply sem_typed_congruence in Hty as (Htycl & Htycl' & Hty); last done. + all: try done. + opose proof* (Hty ∅ ∅ δ_any) as He; first constructor. + revert He. rewrite !subst_map_empty. + simp type_interp. destruct 1 as (v1 & v2 & Hbs1 & Hbs2 & Hv). + simp type_interp in Hv. destruct Hv as (z & -> & ->). eauto. +Qed. + +Lemma soundness_wrt_ctx_equiv Δ Γ e1 e2 A : + TY Δ; Γ ⊢ e1 : A → + TY Δ; Γ ⊢ e2 : A → + TY Δ; Γ ⊨ e1 ≈ e2 : A → + ctx_equiv Δ Γ e1 e2 A. +Proof. + intros ???; eapply sem_typing_ctx_equiv; eauto. + all: by eapply syn_typed_closed. +Qed. diff --git a/theories/type_systems/systemf/exercises05_sol.v b/theories/type_systems/systemf/exercises05_sol.v new file mode 100644 index 0000000..3f7aee1 --- /dev/null +++ b/theories/type_systems/systemf/exercises05_sol.v @@ -0,0 +1,349 @@ +From stdpp Require Import gmap base relations. +From iris Require Import prelude. +From semantics.ts.systemf Require Import lang notation parallel_subst types tactics. +From semantics.ts.systemf Require logrel binary_logrel existential_invariants. + +(** * Exercise Sheet 5 *) + +Implicit Types + (e : expr) + (v : val) + (A B : type) +. + +(** ** Exercise 3 (LN Exercise 23): Existential Fun *) +Section existential. + (** Since extending our language with records would be tedious, + we encode records using nested pairs. + + For instance, we would represent the record type + { add : Int → Int → Int; sub : Int → Int → Int; neg : Int → Int } + as (Int → Int → Int) × (Int → Int → Int) × (Int → Int). + + Similarly, we would represent the record value + { add := λ: "x" "y", "x" + "y"; + sub := λ: "x" "y", "x" - "y"; + neg := λ: "x", #0 - "x" + } + as the nested pair + ((λ: "x" "y", "x" + "y", (* add *) + λ: "x" "y", "x" - "y"), (* sub *) + λ: "x", #0 - "x"). (* neg *) + + *) + + (** We will also assume a recursion combinator. We have not formally added it to our language, but we could do so. *) + Context (Fix : string → string → expr → val). + Notation "'fix:' f x := e" := (Fix f x e)%E + (at level 200, f, x at level 1, e at level 200, + format "'[' 'fix:' f x := '/ ' e ']'") : val_scope. + Notation "'fix:' f x := e" := (Fix f x e)%E + (at level 200, f, x at level 1, e at level 200, + format "'[' 'fix:' f x := '/ ' e ']'") : expr_scope. + Context (fix_typing : ∀ n Γ (f x: string) (A B: type) (e: expr), + type_wf n A → + type_wf n B → + f ≠ x → + TY n; <[x := A]> (<[f := (A → B)%ty]> Γ) ⊢ e : B → + TY n; Γ ⊢ (fix: f x := e) : (A → B)). + + Definition ISET : type :=∃: (#0 (* empty *) + × (Int → #0) (* singleton *) + × (#0 → #0 → #0) (* union *) + × (#0 → #0 → Bool) (* subset *) + ). + + (* We represent sets as functions of type ((Int → Bool) × Int × Int), + storing the mapping, the minimum value, and the maximum value. *) + Definition mini : val := λ: "x" "y", if: "x" < "y" then "x" else "y". + Definition maxi : val := λ: "x" "y", if: "x" < "y" then "x" else "y". + + Definition iterupiv : val := + λ: "f" "max", fix: "rec" "acc" := + let: "i" := Fst "acc" in let: "b" := Snd "acc" in + if: "max" < "i" then "b" else "rec" ("i" + #1, "f" "i" "b"). + Definition iterupv : val := + λ: "f" "init" "min" "max", iterupiv "f" "max" ("min", "init"). + Lemma iterupv_typed n Γ : TY n; Γ ⊢ iterupv : ((Int → Bool → Bool) → Bool → Int → Int → Bool). + Proof. + repeat solve_typing. apply fix_typing; solve_typing. + done. + Qed. + + Definition iset : val :=pack (((λ: "x", #false), #0, #0), (* empty *) + (λ: "n", ((λ: "x", "n" = "x"), "n", "n")), (* singleton *) + (* union *) + (λ: "s1" "s2", ((λ: "x", if: (Fst $ Fst "s1") "x" then #true else (Fst $ Fst "s2") "x"), mini (Snd $ Fst "s1") (Snd $ Fst "s1"), maxi (Snd "s1") (Snd "s1"))), + (* subset *) + (λ: "s1" "s2", + let: "min" := Snd $ Fst "s1" in + let: "max" := Snd $ "s1" in + (* iteration variable is set to #false if we detect a subset violation *) + iterupv (λ: "i" "acc", if: (Fst $ Fst "s2") "i" then "acc" else #false) #true "min" "max" + )). + + Lemma iset_typed n Γ : TY n; Γ ⊢ iset : ISET. + Proof. + (* HINT: use repeated solve_typing with an explicit apply fix_typing inbetween *) + unfold iset. + repeat solve_typing. + by apply fix_typing; solve_typing. + Qed. + + Definition ISETE : type :=∃: (#0 (* empty *) + × (Int → #0) (* singleton *) + × (#0 → #0 → #0) (* union *) + × (#0 → #0 → Bool) (* subset *) + × (#0 → #0 → Bool) (* equality *) + ). + + Definition add_equality : val :=λ: "is", unpack "is" as "isi" in + let: "subset" := Snd "isi" in + pack ("isi", λ: "s1" "s2", if: "subset" "s1" "s2" then "subset" "s2" "s1" else #false). + + Lemma add_equality_typed n Γ : TY n; Γ ⊢ add_equality : (ISET → ISETE)%ty. + Proof. + repeat solve_typing. + + Qed. + +End existential. + +Section ex4. +Import logrel existential_invariants. +(** ** Exercise 4 (LN Exercise 30): Evenness *) +(* Consider the following existential type: *) +Definition even_type : type := + ∃: (#0 × (* zero *) + (#0 → #0) × (* add2 *) + (#0 → Int) (* toint *) + )%ty. + +(* and consider the following implementation of [even_type]: *) +Definition even_impl : val := + pack (#0, + λ: "z", #2 + "z", + λ: "z", "z" + ). +(* We want to prove that [toint] will only every yield even numbers. *) +(* For that purpose, assume that we have a function [even] that decides even parity available: *) +Context (even_dec : val). +Context (even_dec_typed : ∀ n Γ, TY n; Γ ⊢ even_dec : (Int → Bool)). + +(* a) Change [even_impl] to [even_impl_instrumented] such that [toint] asserts evenness of the argument before returned. +You may use the [assert] expression defined in existential_invariants.v. +*) + +Definition even_impl_instrumented : val :=pack (#0, + λ: "z", #2 + "z", + λ: "z", assert (even_dec "z");; "z" + ). + +(* b) Prove that [even_impl_instrumented] is safe. You may assume that even works as intended, + but be sure to state this here. *) +Context (even_spec : ∀ z: Z, big_step (even_dec #z) #(Z.even z)). +Context (even_closed : is_closed [] even_dec). + +Lemma even_impl_instrumented_safe δ: + 𝒱 even_type δ even_impl_instrumented. +Proof. + unfold even_type. simp type_interp. + eexists _. split; first done. + + pose_sem_type (λ v, ∃ z : Z, Z.Even z ∧ v = #z) as τ. + { intros v (z & ? & ->). done. } + exists τ. + + simp type_interp. + eexists _, _. split_and!; first done. + - simp type_interp. eexists _, _. split_and!; first done. + + simp type_interp. simpl. exists 0. split; last done. + apply Z.even_spec. done. + + simp type_interp. eexists _, _. split_and!; [done | done | ]. + intros v. simp type_interp. simpl. + intros (z & Heven & ->). exists #(2 + z)%Z. split; first bs_step_det. + simp type_interp. simpl. exists (2 + z)%Z. + split; last done. destruct Heven as (z' & ->). + exists (z' + 1)%Z. lia. + - simp type_interp. + eexists _, _. split_and!; [done | | ]. + { simpl. rewrite !andb_True. split_and!;[ | done..]. + eapply is_closed_weaken; first done. apply list_subseteq_nil. + } + intros v'. simp type_interp. simpl. intros (z & Heven & ->). + exists #z. + split. + + bs_step_det. eapply bs_if_true; last bs_step_det. + replace true with (Z.even z); first by eapply even_spec. + by apply Z.even_spec. + + simp type_interp. exists z. done. +Qed. +End ex4. + +(** ** Exercise 5 (LN Exercise 31): Abstract sums *) +Section ex5. +Import logrel existential_invariants. +Definition sum_ex_type (A B : type) : type := + ∃: ((A.[ren (+1)] → #0) × + (B.[ren (+1)] → #0) × + (∀: #1 → (A.[ren (+2)] → #0) → (B.[ren (+2)] → #0) → #0) + )%ty. + +Definition sum_ex_impl : val := + pack (λ: "x", (#1, "x"), + λ: "x", (#2, "x"), + Λ, λ: "x" "f1" "f2", if: Fst "x" = #1 then "f1" (Snd "x") else "f2" (Snd "x") + ). + +Lemma sum_ex_safe A B δ: + 𝒱 (sum_ex_type A B) δ sum_ex_impl. +Proof. + intros. unfold sum_ex_type. simp type_interp. + eexists _. split; first done. + + pose_sem_type (λ v, (∃ v', 𝒱 A δ v' ∧ v = (#1, v')%V) ∨ (∃ v', 𝒱 B δ v' ∧ v = (#2, v')%V)) as τ. + { intros v [(v' & Hv & ->) | (v' & Hv & ->)]; simpl; by eapply val_rel_closed. } + exists τ. + + simp type_interp. eexists _, _. split; first done. + split. 1: simp type_interp; eexists _, _; split; first done; split. + - simp type_interp. eexists _, _. split_and!; [done..|]. + intros v' Hv'. simp type_interp. simpl. eexists. split; first bs_step_det. + simp type_interp. simpl. left. + eexists; split; last done. eapply sem_val_rel_cons; done. + - simp type_interp. eexists _, _. split_and!; [done..|]. + intros v' Hv'. simp type_interp. simpl. eexists. split; first bs_step_det. + simp type_interp. simpl. right. + eexists; split; last done. eapply sem_val_rel_cons; done. + - simp type_interp. eexists _. split_and!; [done..|]. + intros τ'. simp type_interp. eexists. split; first bs_step_det. + simp type_interp. eexists _, _; split_and!; [done..|]. + intros v'. simp type_interp. simpl. intros [(v & Hv & ->) | (v & Hv & ->)]. + + eexists. split; first bs_step_det. + specialize (val_rel_closed _ _ _ Hv) as ?. + simp type_interp. eexists _, _; split_and!; [done| simplify_closed | ]. + intros v'. simp type_interp. + intros (? & ? & -> & ? & Hv'). + specialize (Hv' v). simp type_interp in Hv'. destruct Hv' as (? & ? & Hv'). + { revert Hv. rewrite sem_val_rel_cons. rewrite sem_val_rel_cons. asimpl. done. } + eexists _. split; first bs_step_det. + simp type_interp. eexists _, _. split_and!; [done| simplify_closed | ]. + intros v'. simp type_interp. intros (? & ? & -> & ? & _). + eexists _. split. + { bs_step_det. eapply bs_if_true; bs_step_det. + case_decide; bs_step_det. + erewrite lang.subst_is_closed; bs_step_det. + destruct x; simpl; simplify_closed. + } + simp type_interp. simpl. simp type_interp in Hv'. + + eexists. split; first bs_step_det. + specialize (val_rel_closed _ _ _ Hv) as ?. + simp type_interp. eexists _, _; split_and!; [done| simplify_closed | ]. + intros v'. simp type_interp. + intros (? & ? & -> & ? & _). + eexists _. split; first bs_step_det. + simp type_interp. eexists _, _. split_and!; [done| simplify_closed | ]. + intros v'. simp type_interp. intros (? & ? & -> & ? & Hv'). + specialize (Hv' v). simp type_interp in Hv'. destruct Hv' as (? & ? & Hv'). + { revert Hv. rewrite sem_val_rel_cons. rewrite sem_val_rel_cons. asimpl. done. } + + eexists _. split. + { bs_step_det. eapply bs_if_false; bs_step_det. } + simp type_interp. simpl. simp type_interp in Hv'. + Qed. +End ex5. + +(** For Exercise 6 and 7, see binary_logrel.v *) + +(** ** Exercise 8 (LN Exercise 35): Contextual equivalence *) +Section ex8. +Import binary_logrel. +Definition sum_ex_impl' : val := + pack ((λ: "x", InjL "x"), + (λ: "x", InjR "x"), + (Λ, λ: "x" "f1" "f2", Case "x" "f1" "f2") + ). + +Lemma sum_ex_impl'_typed n Γ A B : + type_wf n A → + type_wf n B → + TY n; Γ ⊢ sum_ex_impl' : sum_ex_type A B. +Proof. + intros. + eapply (typed_pack _ _ _ (A + B)%ty). + all: asimpl; solve_typing. +Qed. + +Lemma sum_ex_impl_equiv n Γ A B : + ctx_equiv n Γ sum_ex_impl' sum_ex_impl (sum_ex_type A B). +Proof. + intros. + eapply sem_typing_ctx_equiv; [done | done | ]. + do 2 (split; first done). + intros θ1 θ2 δ Hctx. + rewrite (subst_map_is_closed []); [ | done | intros; simplify_list_elem ]. + rewrite (subst_map_is_closed []); [ | done | intros; simplify_list_elem ]. + simp type_interp. + eexists _, _. split_and!; [bs_step_det.. | ]. + unfold sum_ex_type. simp type_interp. + + eexists _, _; split_and!; [ done | done | ]. + pose_sem_type (λ v1 v2, (∃ v w : val, 𝒱 A δ v w ∧ v1 = InjLV v ∧ v2 = (#1, w)%V) ∨ (∃ v w: val, 𝒱 B δ v w ∧ v1 = InjRV v ∧ v2 = (#2, w)%V)) as R. + { intros ?? [(? & ? & []%val_rel_is_closed & -> & ->) | (? & ? & []%val_rel_is_closed & -> & ->)]; done. } + exists R. + simp type_interp. + eexists _, _, _, _. split; first done. split; first done. split. + - simp type_interp. eexists _, _, _, _. split_and!; [done | done | | ]. + + simp type_interp. eexists _, _, _, _. split_and!; [done | done | done | done | ]. + intros v' w' ?%sem_val_rel_cons. simp type_interp. + simpl. eexists _, _. split_and!; [bs_steps_det | bs_steps_det | ]. + simp type_interp. simpl. left; eauto. + + simp type_interp. eexists _, _, _, _. split_and!; [done | done | done | done | ]. + intros v' w' ?%sem_val_rel_cons. simp type_interp. + simpl. eexists _, _. split_and!; [bs_steps_det | bs_steps_det | ]. + simp type_interp. simpl. right; eauto. + - simp type_interp. eexists _, _. split_and!; [done | done | done | done | ]. + intros R'. simp type_interp. eexists _, _. split_and!; [bs_steps_det | bs_steps_det | ]. + simp type_interp. eexists _, _, _, _. split_and!; [ done | done | done | done | ]. + intros v' w'. simp type_interp. simpl. intros Hsum. + assert (is_closed [] v' ∧ is_closed [] w') as [Hclv' Hclw']. + { destruct Hsum as [(? & ? & []%val_rel_is_closed & -> & ->) | (? & ? & []%val_rel_is_closed & -> & ->)]; done. } + eexists _, _. split_and!; [bs_steps_det | bs_steps_det | ]. + simp type_interp. eexists _, _, _, _. + split_and!; [ done | done | simplify_closed | simplify_closed | ]. + intros f f' Hf. + simpl. repeat (rewrite subst_is_closed_nil; [ | done]). + simp type_interp. eexists _, _. + split_and!; [ bs_steps_det | bs_steps_det | ]. + simp type_interp. eexists _, _, _, _. + specialize (val_rel_is_closed _ _ _ _ Hf) as []. + split_and!; [done | done | simplify_closed | simplify_closed | ]. + intros g g' Hg. + simpl. repeat (rewrite subst_is_closed_nil; [ | done]). + (* CA *) + destruct Hsum as [(v & w & Hvw & -> & ->) | (v & w & Hvw & -> & ->)]. + + simpl; simp type_interp. + simp type_interp in Hf. destruct Hf as (? & ? & ? & ? & -> & -> & ? & ? & Hf). + opose proof* (Hf v w) as Hf. + { revert Hvw. rewrite sem_val_rel_cons. rewrite sem_val_rel_cons. asimpl. done. } + simp type_interp in Hf. destruct Hf as (v1 & v2 & ? & ? & Hf). + simp type_interp in Hf. simpl in Hf. + eexists _, _. + split_and!. + { eapply bs_casel; bs_step_det. } + { eapply bs_if_true; bs_step_det. } + done. + + simpl; simp type_interp. + simp type_interp in Hg. destruct Hg as (? & ? & ? & ? & -> & -> & ? & ? & Hg). + opose proof* (Hg v w) as Hg. + { revert Hvw. rewrite sem_val_rel_cons. rewrite sem_val_rel_cons. asimpl. done. } + simp type_interp in Hg. destruct Hg as (v1 & v2 & ? & ? & Hg). + simp type_interp in Hg. simpl in Hg. + eexists _, _. + split_and!. + { eapply bs_caser; bs_step_det. } + { eapply bs_if_false; bs_step_det. } + done. +Qed. + +End ex8. diff --git a/theories/type_systems/systemf/logrel_sol.v b/theories/type_systems/systemf/logrel_sol.v index 1d21bbc..7b97b2f 100644 --- a/theories/type_systems/systemf/logrel_sol.v +++ b/theories/type_systems/systemf/logrel_sol.v @@ -583,18 +583,47 @@ Lemma compat_pack Δ Γ e n A B : TY n; Γ ⊨ e : A.[B/] → TY n; Γ ⊨ (pack e) : (∃: A). Proof. -(* This will be an exercise for you next week :) *) -Admitted. + (* This will be an exercise for you next week :) *) + intros Hwf1 Hwf2 [Hecl He]. + split; first naive_solver. + intros θ δ Hctx. simpl. + simp type_interp. + + specialize (He _ _ Hctx). simp type_interp in He. + destruct He as (v & Hb & Hv). + exists (PackV v). split; first eauto. + simp type_interp. exists v. split; first done. + exists (interp_type B δ). + apply sem_val_rel_move_single_subst. done. +Qed. Lemma compat_unpack n Γ A B e e' x : -type_wf n B → -TY n; Γ ⊨ e : (∃: A) → -TY S n; <[x:=A]> (⤉Γ) ⊨ e' : B.[ren (+1)] → -TY n; Γ ⊨ (unpack e as BNamed x in e') : B. + type_wf n B → + TY n; Γ ⊨ e : (∃: A) → + TY S n; <[x:=A]> (⤉Γ) ⊨ e' : B.[ren (+1)] → + TY n; Γ ⊨ (unpack e as BNamed x in e') : B. Proof. -(* This will be an exercise for you next week :) *) -Admitted. + (* This will be an exercise for you next week :) *) + intros Hwf [Hecl He] [He'cl He']. split. + { simpl. apply andb_True. split; first done. eapply is_closed_weaken; first eassumption. set_solver. } + intros θ δ Hctx. simpl. + simp type_interp. + + specialize (He _ _ Hctx). simp type_interp in He. + destruct He as (v & Hb & Hv). + simp type_interp in Hv. destruct Hv as (v' & -> & τ & Hv'). + + specialize (He' (<[x := of_val v']> θ) (τ.:δ)). + simp type_interp in He'. + destruct He' as (v & Hb' & Hv). + { constructor; first done. by apply sem_context_rel_cons. } + exists v. split. + { econstructor; first done. rewrite subst'_subst_map; first done. + by eapply sem_context_rel_closed. + } + by eapply sem_val_rel_cons. +Qed. Lemma compat_if n Γ e0 e1 e2 A : TY n; Γ ⊨ e0 : Bool → diff --git a/theories/type_systems/systemf/types_sol.v b/theories/type_systems/systemf/types_sol.v new file mode 100644 index 0000000..b3703e5 --- /dev/null +++ b/theories/type_systems/systemf/types_sol.v @@ -0,0 +1,875 @@ +From stdpp Require Import base relations. +From iris Require Import prelude. +From semantics.lib Require Import maps. +From semantics.ts.systemf Require Import lang notation. +From Autosubst Require Export Autosubst. + +(** ** Syntactic typing *) +(** We use De Bruijn indices with the help of the Autosubst library. *) +Inductive type : Type := + (** [var] is the type of variables of Autosubst -- it unfolds to [nat] *) + | TVar : var → type + | Int + | Bool + | Unit + (** The [{bind 1 of type}] tells Autosubst to put a De Bruijn binder here *) + | TForall : {bind 1 of type} → type + | TExists : {bind 1 of type} → type + | Fun (A B : type) + | Prod (A B : type) + | Sum (A B : type). + +(** Autosubst instances. + This lets Autosubst do its magic and derive all the substitution functions, etc. + *) +#[export] Instance Ids_type : Ids type. derive. Defined. +#[export] Instance Rename_type : Rename type. derive. Defined. +#[export] Instance Subst_type : Subst type. derive. Defined. +#[export] Instance SubstLemmas_typer : SubstLemmas type. derive. Qed. + +Definition typing_context := gmap string type. +Implicit Types + (Γ : typing_context) + (v : val) + (e : expr). + +Declare Scope FType_scope. +Delimit Scope FType_scope with ty. +Bind Scope FType_scope with type. +Notation "# x" := (TVar x) : FType_scope. +Infix "→" := Fun : FType_scope. +Notation "(→)" := Fun (only parsing) : FType_scope. +Notation "∀: τ" := + (TForall τ%ty) + (at level 100, τ at level 200) : FType_scope. +Notation "∃: τ" := + (TExists τ%ty) + (at level 100, τ at level 200) : FType_scope. +Infix "×" := Prod (at level 70) : FType_scope. +Notation "(×)" := Prod (only parsing) : FType_scope. +Infix "+" := Sum : FType_scope. +Notation "(+)" := Sum (only parsing) : FType_scope. + + +(** Shift all the indices in the context by one, + used when inserting a new type interpretation in Δ. *) +(* [<$>] is notation for the [fmap] operation that maps the substitution over the whole map. *) +(* [ren] is Autosubst's renaming operation -- it renames all type variables according to the given function, + in this case [(+1)] to shift the variables up by 1. *) +Notation "⤉ Γ" := (Autosubst_Classes.subst (ren (+1)) <$> Γ) (at level 10, format "⤉ Γ"). + +Reserved Notation "'TY' n ; Γ ⊢ e : A" (at level 74, e, A at next level). + +(** [type_wf n A] states that a type [A] has only free variables up to < [n]. + (in other words, all variables occurring free are strictly bounded by [n]). *) +Inductive type_wf : nat → type → Prop := + | type_wf_TVar m n: + m < n → + type_wf n (TVar m) + | type_wf_Int n: type_wf n Int + | type_wf_Bool n : type_wf n Bool + | type_wf_Unit n : type_wf n Unit + | type_wf_TForall n A : + type_wf (S n) A → + type_wf n (TForall A) + | type_wf_TExists n A : + type_wf (S n) A → + type_wf n (TExists A) + | type_wf_Fun n A B: + type_wf n A → + type_wf n B → + type_wf n (Fun A B) + | type_wf_Prod n A B : + type_wf n A → + type_wf n B → + type_wf n (Prod A B) + | type_wf_Sum n A B : + type_wf n A → + type_wf n B → + type_wf n (Sum A B) +. +#[export] Hint Constructors type_wf : core. + +Inductive bin_op_typed : bin_op → type → type → type → Prop := + | plus_op_typed : bin_op_typed PlusOp Int Int Int + | minus_op_typed : bin_op_typed MinusOp Int Int Int + | mul_op_typed : bin_op_typed MultOp Int Int Int + | lt_op_typed : bin_op_typed LtOp Int Int Bool + | le_op_typed : bin_op_typed LeOp Int Int Bool + | eq_op_typed : bin_op_typed EqOp Int Int Bool. +#[export] Hint Constructors bin_op_typed : core. + +Inductive un_op_typed : un_op → type → type → Prop := + | neg_op_typed : un_op_typed NegOp Bool Bool + | minus_un_op_typed : un_op_typed MinusUnOp Int Int. + +Inductive syn_typed : nat → typing_context → expr → type → Prop := + | typed_var n Γ x A : + Γ !! x = Some A → + TY n; Γ ⊢ (Var x) : A + | typed_lam n Γ x e A B : + TY n ; (<[ x := A]> Γ) ⊢ e : B → + type_wf n A → + TY n; Γ ⊢ (Lam (BNamed x) e) : (A → B) + | typed_lam_anon n Γ e A B : + TY n ; Γ ⊢ e : B → + type_wf n A → + TY n; Γ ⊢ (Lam BAnon e) : (A → B) + | typed_tlam n Γ e A : + (* we need to shift the context up as we descend under a binder *) + TY S n; (⤉ Γ) ⊢ e : A → + TY n; Γ ⊢ (Λ, e) : (∀: A) + | typed_tapp n Γ A B e : + TY n; Γ ⊢ e : (∀: A) → + type_wf n B → + (* A.[B/] is the notation for Autosubst's substitution operation that + replaces variable 0 by [B] *) + TY n; Γ ⊢ (e <>) : (A.[B/]) + | typed_pack n Γ A B e : + type_wf n B → + type_wf (S n) A → + TY n; Γ ⊢ e : (A.[B/]) → + TY n; Γ ⊢ (pack e) : (∃: A) + | typed_unpack n Γ A B e e' x : + type_wf n B → (* we should not leak the existential! *) + TY n; Γ ⊢ e : (∃: A) → + (* As we descend under a type variable binder for the typing of [e'], + we need to shift the indices in [Γ] and [B] up by one. + On the other hand, [A] is already defined under this binder, so we need not shift it. + *) + TY (S n); (<[x := A]>(⤉Γ)) ⊢ e' : (B.[ren (+1)]) → + TY n; Γ ⊢ (unpack e as BNamed x in e') : B + | typed_int n Γ z : TY n; Γ ⊢ (Lit $ LitInt z) : Int + | typed_bool n Γ b : TY n; Γ ⊢ (Lit $ LitBool b) : Bool + | typed_unit n Γ : TY n; Γ ⊢ (Lit $ LitUnit) : Unit + | typed_if n Γ e0 e1 e2 A : + TY n; Γ ⊢ e0 : Bool → + TY n; Γ ⊢ e1 : A → + TY n; Γ ⊢ e2 : A → + TY n; Γ ⊢ If e0 e1 e2 : A + | typed_app n Γ e1 e2 A B : + TY n; Γ ⊢ e1 : (A → B) → + TY n; Γ ⊢ e2 : A → + TY n; Γ ⊢ (e1 e2)%E : B + | typed_binop n Γ e1 e2 op A B C : + bin_op_typed op A B C → + TY n; Γ ⊢ e1 : A → + TY n; Γ ⊢ e2 : B → + TY n; Γ ⊢ BinOp op e1 e2 : C + | typed_unop n Γ e op A B : + un_op_typed op A B → + TY n; Γ ⊢ e : A → + TY n; Γ ⊢ UnOp op e : B + | typed_pair n Γ e1 e2 A B : + TY n; Γ ⊢ e1 : A → + TY n; Γ ⊢ e2 : B → + TY n; Γ ⊢ (e1, e2) : A × B + | typed_fst n Γ e A B : + TY n; Γ ⊢ e : A × B → + TY n; Γ ⊢ Fst e : A + | typed_snd n Γ e A B : + TY n; Γ ⊢ e : A × B → + TY n; Γ ⊢ Snd e : B + | typed_injl n Γ e A B : + type_wf n B → + TY n; Γ ⊢ e : A → + TY n; Γ ⊢ InjL e : A + B + | typed_injr n Γ e A B : + type_wf n A → + TY n; Γ ⊢ e : B → + TY n; Γ ⊢ InjR e : A + B + | typed_case n Γ e e1 e2 A B C : + TY n; Γ ⊢ e : B + C → + TY n; Γ ⊢ e1 : (B → A) → + TY n; Γ ⊢ e2 : (C → A) → + TY n; Γ ⊢ Case e e1 e2 : A +where "'TY' n ; Γ ⊢ e : A" := (syn_typed n Γ e%E A%ty). +#[export] Hint Constructors syn_typed : core. + +(** Examples *) +Goal TY 0; ∅ ⊢ (λ: "x", #1 + "x")%E : (Int → Int). +Proof. eauto. Qed. +(** [∀: #0 → #0] corresponds to [∀ α. α → α] with named binders. *) +Goal TY 0; ∅ ⊢ (Λ, λ: "x", "x")%E : (∀: #0 → #0). +Proof. repeat econstructor. Qed. +Goal TY 0; ∅ ⊢ (pack ((λ: "x", "x"), #42)) : ∃: (#0 → #0) × #0. +Proof. + apply (typed_pack _ _ _ Int). + - eauto. + - repeat econstructor. + - (* [asimpl] is Autosubst's tactic for simplifying goals involving type substitutions. *) + asimpl. eauto. +Qed. +Goal TY 0; ∅ ⊢ (unpack (pack ((λ: "x", "x"), #42)) as "y" in (λ: "x", #1337) ((Fst "y") (Snd "y"))) : Int. +Proof. + (* if we want to typecheck stuff with existentials, we need a bit more explicit proofs. + Letting eauto try to instantiate the evars becomes too expensive. *) + apply (typed_unpack _ _ ((#0 → #0) × #0)%ty). + - done. + - apply (typed_pack _ _ _ Int); asimpl; eauto. + repeat econstructor. + - eapply (typed_app _ _ _ _ (#0)%ty); eauto 10. +Qed. + +(** fails: we are not allowed to leak the existential *) +Goal TY 0; ∅ ⊢ (unpack (pack ((λ: "x", "x"), #42)) as "y" in (Fst "y") (Snd "y")) : #0. +Proof. + apply (typed_unpack _ _ ((#0 → #0) × #0)%ty). +Abort. + +(* derived typing rule for match *) +Lemma typed_match n Γ e e1 e2 x1 x2 A B C : + type_wf n B → + type_wf n C → + TY n; Γ ⊢ e : B + C → + TY n; <[x1 := B]> Γ ⊢ e1 : A → + TY n; <[x2 := C]> Γ ⊢ e2 : A → + TY n; Γ ⊢ match: e with InjL (BNamed x1) => e1 | InjR (BNamed x2) => e2 end : A. +Proof. eauto. Qed. + +Lemma syn_typed_closed n Γ e A X : + TY n ; Γ ⊢ e : A → + (∀ x, x ∈ dom Γ → x ∈ X) → + is_closed X e. +Proof. + induction 1 as [ | ??????? IH | | n Γ e A H IH | | | n Γ A B e e' x Hwf H1 IH1 H2 IH2 | | | | | | | | | | | | | ] in X |-*; simpl; intros Hx; try done. + + { (* var *) apply bool_decide_pack, Hx. apply elem_of_dom; eauto. } + { (* lam *) apply IH. + intros y. rewrite elem_of_dom lookup_insert_is_Some. + intros [<- | [? Hy]]; first by apply elem_of_cons; eauto. + apply elem_of_cons. right. eapply Hx. by apply elem_of_dom. + } + { (* anon lam *) naive_solver. } + { (* tlam *) + eapply IH. intros x Hel. apply Hx. + by rewrite dom_fmap in Hel. + } + 3: { (* unpack *) + apply andb_True; split. + - apply IH1. apply Hx. + - apply IH2. intros y. rewrite elem_of_dom lookup_insert_is_Some. + intros [<- | [? Hy]]; first by apply elem_of_cons; eauto. + apply elem_of_cons. right. eapply Hx. + apply elem_of_dom. revert Hy. rewrite lookup_fmap fmap_is_Some. done. + } + (* everything else *) + all: repeat match goal with + | |- Is_true (_ && _) => apply andb_True; split + end. + all: try naive_solver. +Qed. + +(** *** Lemmas about [type_wf] *) +Lemma type_wf_mono n m A: + type_wf n A → n ≤ m → type_wf m A. +Proof. + induction 1 in m |-*; eauto with lia. +Qed. + +Lemma type_wf_rename n A δ: + type_wf n A → + (∀ i j, i < j → δ i < δ j) → + type_wf (δ n) (rename δ A). +Proof. + induction 1 in δ |-*; intros Hmon; simpl; eauto. + all: econstructor; eapply type_wf_mono; first eapply IHtype_wf; last done. + all: intros i j Hlt; destruct i, j; simpl; try lia. + all: rewrite -Nat.succ_lt_mono; eapply Hmon; lia. +Qed. + +(** [A.[σ]], i.e. [A] with the substitution [σ] applied to it, is well-formed under [m] if + [A] is well-formed under [n] and all the things we substitute up to [n] are well-formed under [m]. + *) +Lemma type_wf_subst n m A σ: + type_wf n A → + (∀ x, x < n → type_wf m (σ x)) → + type_wf m A.[σ]. +Proof. + induction 1 in m, σ |-*; intros Hsub; simpl; eauto. + + econstructor; eapply IHtype_wf. + intros [|x]; rewrite /up //=. + - econstructor. lia. + - intros Hlt % Nat.succ_lt_mono. eapply type_wf_rename; eauto. + intros i j Hlt'; simpl; lia. + + econstructor; eapply IHtype_wf. + intros [|x]; rewrite /up //=. + - econstructor. lia. + - intros Hlt % Nat.succ_lt_mono. eapply type_wf_rename; eauto. + intros i j Hlt'; simpl; lia. +Qed. + +Lemma type_wf_single_subst n A B: type_wf n B → type_wf (S n) A → type_wf n A.[B/]. +Proof. + intros HB HA. eapply type_wf_subst; first done. + intros [|x]; simpl; eauto. + intros ?; econstructor. lia. +Qed. + +(** We lift [type_wf] to well-formedness of contexts *) +Definition ctx_wf n Γ := (∀ x A, Γ !! x = Some A → type_wf n A). + +Lemma ctx_wf_empty n : ctx_wf n ∅. +Proof. rewrite /ctx_wf. set_solver. Qed. + +Lemma ctx_wf_insert n x Γ A: ctx_wf n Γ → type_wf n A → ctx_wf n (<[x := A]> Γ). +Proof. intros H1 H2 y B. rewrite lookup_insert_Some. naive_solver. Qed. + +Lemma ctx_wf_up n Γ: + ctx_wf n Γ → ctx_wf (S n) (⤉Γ). +Proof. + intros Hwf x A; rewrite lookup_fmap. + intros (B & Hlook & ->) % fmap_Some. + asimpl. eapply type_wf_subst; first eauto. + intros y Hlt. simpl. econstructor. lia. +Qed. + +#[global] +Hint Resolve ctx_wf_empty ctx_wf_insert ctx_wf_up : core. + +(** Well-typed terms at [A] under a well-formed context have well-formed types [A].*) +Lemma syn_typed_wf n Γ e A: + ctx_wf n Γ → + TY n; Γ ⊢ e : A → + type_wf n A. +Proof. + intros Hwf; induction 1 as [ | n Γ x e A B Hty IH Hwfty | | n Γ e A Hty IH | n Γ A B e Hty IH Hwfty | n Γ A B e Hwfty Hty IH| | | | | | n Γ e1 e2 A B HtyA IHA HtyB IHB | n Γ e1 e2 op A B C Hop HtyA IHA HtyB IHB | n Γ e op A B Hop H IH | n Γ e1 e2 A B HtyA IHA HtyB IHB | n Γ e A B Hty IH | n Γ e A B Hty IH | n Γ e A B Hwfty Hty IH | n Γ e A B Hwfty Hty IH| n Γ e e1 e2 A B C Htye IHe Htye1 IHe1 Htye2 IHe2 ]; eauto. + - eapply type_wf_single_subst; first done. + specialize (IH Hwf) as Hwf'. + by inversion Hwf'. + - specialize (IHA Hwf) as Hwf'. + by inversion Hwf'; subst. + - inversion Hop; subst; eauto. + - inversion Hop; subst; eauto. + - specialize (IH Hwf) as Hwf'. by inversion Hwf'; subst. + - specialize (IH Hwf) as Hwf'. by inversion Hwf'; subst. + - specialize (IHe1 Hwf) as Hwf''. by inversion Hwf''; subst. +Qed. + +Lemma renaming_inclusion Γ Δ : Γ ⊆ Δ → ⤉Γ ⊆ ⤉Δ. +Proof. + eapply map_fmap_mono. +Qed. + +Lemma typed_weakening n m Γ Δ e A: + TY n; Γ ⊢ e : A → + Γ ⊆ Δ → + n ≤ m → + TY m; Δ ⊢ e : A. +Proof. + induction 1 as [| n Γ x e A B Htyp IH | | n Γ e A Htyp IH | | |n Γ A B e e' x Hwf H1 IH1 H2 IH2 | | | | | | | | | | | | | ] in Δ, m |-*; intros Hsub Hle; eauto using type_wf_mono. + - (* var *) econstructor. by eapply lookup_weaken. + - (* lam *) econstructor; last by eapply type_wf_mono. eapply IH; eauto. by eapply insert_mono. + - (* tlam *) econstructor. eapply IH; last by lia. by eapply renaming_inclusion. + - (* pack *) + econstructor; last naive_solver. all: (eapply type_wf_mono; [ done | lia]). + - (* unpack *) econstructor. + + eapply type_wf_mono; done. + + eapply IH1; done. + + eapply IH2; last lia. apply insert_mono. by apply renaming_inclusion. +Qed. + +Lemma type_wf_subst_dom σ τ n A: + type_wf n A → + (∀ m, m < n → σ m = τ m) → + A.[σ] = A.[τ]. +Proof. + induction 1 in σ, τ |-*; simpl; eauto. + - (* tforall *) + intros Heq; asimpl. f_equal. + eapply IHtype_wf; intros [|m]; rewrite /up; simpl; first done. + intros Hlt. f_equal. eapply Heq. lia. + - (* texists *) + intros Heq; asimpl. f_equal. + eapply IHtype_wf. intros [ | m]; rewrite /up; simpl; first done. + intros Hlt. f_equal. apply Heq. lia. + - (* fun *) intros ?. f_equal; eauto. + - (* prod *) intros ?. f_equal; eauto. + - (* sum *) intros ?. f_equal; eauto. +Qed. + +Lemma type_wf_closed A σ: + type_wf 0 A → + A.[σ] = A. +Proof. + intros Hwf; erewrite (type_wf_subst_dom _ (ids) 0). + - by asimpl. + - done. + - intros ??; lia. +Qed. + +(** Typing inversion lemmas *) +Lemma var_inversion Γ n (x: string) A: TY n; Γ ⊢ x : A → Γ !! x = Some A. +Proof. inversion 1; subst; auto. Qed. + +Lemma lam_inversion n Γ (x: string) e C: + TY n; Γ ⊢ (λ: x, e) : C → + ∃ A B, C = (A → B)%ty ∧ type_wf n A ∧ TY n; <[x:=A]> Γ ⊢ e : B. +Proof. inversion 1; subst; eauto 10. Qed. + +Lemma lam_anon_inversion n Γ e C: + TY n; Γ ⊢ (λ: <>, e) : C → + ∃ A B, C = (A → B)%ty ∧ type_wf n A ∧ TY n; Γ ⊢ e : B. +Proof. inversion 1; subst; eauto 10. Qed. + +Lemma app_inversion n Γ e1 e2 B: + TY n; Γ ⊢ e1 e2 : B → + ∃ A, TY n; Γ ⊢ e1 : (A → B) ∧ TY n; Γ ⊢ e2 : A. +Proof. inversion 1; subst; eauto. Qed. + +Lemma if_inversion n Γ e0 e1 e2 B: + TY n; Γ ⊢ If e0 e1 e2 : B → + TY n; Γ ⊢ e0 : Bool ∧ TY n; Γ ⊢ e1 : B ∧ TY n; Γ ⊢ e2 : B. +Proof. inversion 1; subst; eauto. Qed. + +Lemma binop_inversion n Γ op e1 e2 B: + TY n; Γ ⊢ BinOp op e1 e2 : B → + ∃ A1 A2, bin_op_typed op A1 A2 B ∧ TY n; Γ ⊢ e1 : A1 ∧ TY n; Γ ⊢ e2 : A2. +Proof. inversion 1; subst; eauto. Qed. + +Lemma unop_inversion n Γ op e B: + TY n; Γ ⊢ UnOp op e : B → + ∃ A, un_op_typed op A B ∧ TY n; Γ ⊢ e : A. +Proof. inversion 1; subst; eauto. Qed. + +Lemma type_app_inversion n Γ e B: + TY n; Γ ⊢ e <> : B → + ∃ A C, B = A.[C/] ∧ type_wf n C ∧ TY n; Γ ⊢ e : (∀: A). +Proof. inversion 1; subst; eauto. Qed. + +Lemma type_lam_inversion n Γ e B: + TY n; Γ ⊢ (Λ,e) : B → + ∃ A, B = (∀: A)%ty ∧ TY (S n); ⤉Γ ⊢ e : A. +Proof. inversion 1; subst; eauto. Qed. + +Lemma type_pack_inversion n Γ e B : + TY n; Γ ⊢ (pack e) : B → + ∃ A C, B = (∃: A)%ty ∧ TY n; Γ ⊢ e : (A.[C/])%ty ∧ type_wf n C ∧ type_wf (S n) A. +Proof. inversion 1; subst; eauto 10. Qed. + +Lemma type_unpack_inversion n Γ e e' x B : + TY n; Γ ⊢ (unpack e as x in e') : B → + ∃ A x', x = BNamed x' ∧ type_wf n B ∧ TY n; Γ ⊢ e : (∃: A) ∧ TY S n; <[x' := A]> (⤉Γ) ⊢ e' : (B.[ren (+1)]). +Proof. inversion 1; subst; eauto 10. Qed. + +Lemma pair_inversion n Γ e1 e2 C : + TY n; Γ ⊢ (e1, e2) : C → + ∃ A B, C = (A × B)%ty ∧ TY n; Γ ⊢ e1 : A ∧ TY n; Γ ⊢ e2 : B. +Proof. inversion 1; subst; eauto. Qed. + +Lemma fst_inversion n Γ e A : + TY n; Γ ⊢ Fst e : A → + ∃ B, TY n; Γ ⊢ e : A × B. +Proof. inversion 1; subst; eauto. Qed. + +Lemma snd_inversion n Γ e B : + TY n; Γ ⊢ Snd e : B → + ∃ A, TY n; Γ ⊢ e : A × B. +Proof. inversion 1; subst; eauto. Qed. + +Lemma injl_inversion n Γ e C : + TY n; Γ ⊢ InjL e : C → + ∃ A B, C = (A + B)%ty ∧ TY n; Γ ⊢ e : A ∧ type_wf n B. +Proof. inversion 1; subst; eauto. Qed. + +Lemma injr_inversion n Γ e C : + TY n; Γ ⊢ InjR e : C → + ∃ A B, C = (A + B)%ty ∧ TY n; Γ ⊢ e : B ∧ type_wf n A. +Proof. inversion 1; subst; eauto. Qed. + +Lemma case_inversion n Γ e e1 e2 A : + TY n; Γ ⊢ Case e e1 e2 : A → + ∃ B C, TY n; Γ ⊢ e : B + C ∧ TY n; Γ ⊢ e1 : (B → A) ∧ TY n; Γ ⊢ e2 : (C → A). +Proof. inversion 1; subst; eauto. Qed. + + +Lemma typed_substitutivity n e e' Γ (x: string) A B : + TY 0; ∅ ⊢ e' : A → + TY n; (<[x := A]> Γ) ⊢ e : B → + TY n; Γ ⊢ lang.subst x e' e : B. +Proof. + intros He'. revert n B Γ; induction e as [| y | y | | | | | | | | | | | | | | ]; intros n B Γ; simpl. + - inversion 1; subst; auto. + - intros Hp % var_inversion. + destruct (decide (x = y)). + + subst. rewrite lookup_insert in Hp. injection Hp as ->. + eapply typed_weakening; [done| |lia]. apply map_empty_subseteq. + + rewrite lookup_insert_ne in Hp; last done. auto. + - destruct y as [ | y]. + { intros (A' & C & -> & Hwf & Hty) % lam_anon_inversion. + econstructor; last done. destruct decide as [Heq|]. + + congruence. + + eauto. + } + intros (A' & C & -> & Hwf & Hty) % lam_inversion. + econstructor; last done. destruct decide as [Heq|]. + + injection Heq as [= ->]. by rewrite insert_insert in Hty. + + rewrite insert_commute in Hty; last naive_solver. eauto. + - intros (C & Hty1 & Hty2) % app_inversion. eauto. + - intros (? & Hop & H1) % unop_inversion. + destruct op; inversion Hop; subst; eauto. + - intros (? & ? & Hop & H1 & H2) % binop_inversion. + destruct op; inversion Hop; subst; eauto. + - intros (H1 & H2 & H3)%if_inversion. naive_solver. + - intros (C & D & -> & Hwf & Hty) % type_app_inversion. eauto. + - intros (C & -> & Hty)%type_lam_inversion. econstructor. + eapply IHe. revert Hty. rewrite fmap_insert. + eapply syn_typed_wf in He'; last by naive_solver. + rewrite type_wf_closed; eauto. + - intros (C & D & -> & Hty & Hwf1 & Hwf2)%type_pack_inversion. + econstructor; [done..|]. apply IHe. done. + - intros (C & x' & -> & Hwf & Hty1 & Hty2)%type_unpack_inversion. + econstructor; first done. + + eapply IHe1. done. + + destruct decide as [Heq | ]. + * injection Heq as [= ->]. by rewrite fmap_insert insert_insert in Hty2. + * rewrite fmap_insert in Hty2. rewrite insert_commute in Hty2; last naive_solver. + eapply IHe2. rewrite type_wf_closed in Hty2; first done. + eapply syn_typed_wf; last apply He'. done. + - intros (? & ? & -> & ? & ?) % pair_inversion. eauto. + - intros (? & ?)%fst_inversion. eauto. + - intros (? & ?)%snd_inversion. eauto. + - intros (? & ? & -> & ? & ?)%injl_inversion. eauto. + - intros (? & ? & -> & ? & ?)%injr_inversion. eauto. + - intros (? & ? & ? & ? & ?)%case_inversion. eauto. +Qed. + +(** Canonical values *) +Lemma canonical_values_arr n Γ e A B: + TY n; Γ ⊢ e : (A → B) → + is_val e → + ∃ x e', e = (λ: x, e')%E. +Proof. + inversion 1; simpl; naive_solver. +Qed. + +Lemma canonical_values_forall n Γ e A: + TY n; Γ ⊢ e : (∀: A)%ty → + is_val e → + ∃ e', e = (Λ, e')%E. +Proof. + inversion 1; simpl; naive_solver. +Qed. + +Lemma canonical_values_exists n Γ e A : + TY n; Γ ⊢ e : (∃: A) → + is_val e → + ∃ e', e = (pack e')%E. +Proof. inversion 1; simpl; naive_solver. Qed. + +Lemma canonical_values_int n Γ e: + TY n; Γ ⊢ e : Int → + is_val e → + ∃ n: Z, e = (#n)%E. +Proof. + inversion 1; simpl; naive_solver. +Qed. + +Lemma canonical_values_bool n Γ e: + TY n; Γ ⊢ e : Bool → + is_val e → + ∃ b: bool, e = (#b)%E. +Proof. + inversion 1; simpl; naive_solver. +Qed. + +Lemma canonical_values_unit n Γ e: + TY n; Γ ⊢ e : Unit → + is_val e → + e = (#LitUnit)%E. +Proof. + inversion 1; simpl; naive_solver. +Qed. + +Lemma canonical_values_prod n Γ e A B : + TY n; Γ ⊢ e : A × B → + is_val e → + ∃ e1 e2, e = (e1, e2)%E ∧ is_val e1 ∧ is_val e2. +Proof. + inversion 1; simpl; naive_solver. +Qed. + +Lemma canonical_values_sum n Γ e A B : + TY n; Γ ⊢ e : A + B → + is_val e → + (∃ e', e = InjL e' ∧ is_val e') ∨ (∃ e', e = InjR e' ∧ is_val e'). +Proof. + inversion 1; simpl; naive_solver. +Qed. + +(** Progress *) +Lemma typed_progress e A: + TY 0; ∅ ⊢ e : A → is_val e ∨ reducible e. +Proof. + remember ∅ as Γ. remember 0 as n. + induction 1 as [| | | | n Γ A B e Hty IH | n Γ A B e Hwf Hwf' Hty IH | n Γ A B e e' x Hwf Hty1 IH1 Hty2 IH2 | | | | n Γ e0 e1 e2 A Hty1 IH1 Hty2 IH2 Hty3 IH3 | n Γ e1 e2 A B Hty IH1 _ IH2 | n Γ e1 e2 op A B C Hop Hty1 IH1 Hty2 IH2 | n Γ e op A B Hop Hty IH | n Γ e1 e2 A B Hty1 IH1 Hty2 IH2 | n Γ e A B Hty IH | n Γ e A B Hty IH | n Γ e A B Hwf Hty IH | n Γ e A B Hwf Hty IH| n Γ e e1 e2 A B C Htye IHe Htye1 IHe1 Htye2 IHe2]. + - subst. naive_solver. + - left. done. + - left. done. + - (* big lambda *) left; done. + - (* type app *) + right. destruct (IH HeqΓ Heqn) as [H1|H1]. + + eapply canonical_values_forall in Hty as [e' ->]; last done. + eexists. eapply base_contextual_step. eapply TBetaS. + + destruct H1 as [e' H1]. eexists. eauto. + - (* pack *) + destruct (IH HeqΓ Heqn) as [H | H]. + + by left. + + right. destruct H as [e' H]. eexists. eauto. + - (* unpack *) + destruct (IH1 HeqΓ Heqn) as [H | H]. + + eapply canonical_values_exists in Hty1 as [e'' ->]; last done. + right. eexists. eapply base_contextual_step. constructor; last done. + done. + + right. destruct H as [e'' H]. eexists. eauto. + - (* int *)left. done. + - (* bool*) left. done. + - (* unit *) left. done. + - (* if *) + destruct (IH1 HeqΓ Heqn) as [H1 | H1]. + + eapply canonical_values_bool in Hty1 as (b & ->); last done. + right. destruct b; eexists; eapply base_contextual_step; constructor. + + right. destruct H1 as [e0' Hstep]. + eexists. eauto. + - (* app *) + destruct (IH2 HeqΓ Heqn) as [H2|H2]; [destruct (IH1 HeqΓ Heqn) as [H1|H1]|]. + + eapply canonical_values_arr in Hty as (x & e & ->); last done. + right. eexists. + eapply base_contextual_step, BetaS; eauto. + + right. destruct H1 as [e1' Hstep]. + eexists. eauto. + + right. destruct H2 as [e2' H2]. + eexists. eauto. + - (* binop *) + assert (A = Int ∧ B = Int) as [-> ->]. + { inversion Hop; subst A B C; done. } + destruct (IH2 HeqΓ Heqn) as [H2|H2]; [destruct (IH1 HeqΓ Heqn) as [H1|H1]|]. + + right. eapply canonical_values_int in Hty1 as [n1 ->]; last done. + eapply canonical_values_int in Hty2 as [n2 ->]; last done. + inversion Hop; subst; simpl. + all: eexists; eapply base_contextual_step; eapply BinOpS; eauto. + + right. destruct H1 as [e1' Hstep]. eexists. eauto. + + right. destruct H2 as [e2' H2]. eexists. eauto. + - (* unop *) + inversion Hop; subst A B op. + + right. destruct (IH HeqΓ Heqn) as [H2 | H2]. + * eapply canonical_values_bool in Hty as [b ->]; last done. + eexists; eapply base_contextual_step; eapply UnOpS; eauto. + * destruct H2 as [e' H2]. eexists. eauto. + + right. destruct (IH HeqΓ Heqn) as [H2 | H2]. + * eapply canonical_values_int in Hty as [z ->]; last done. + eexists; eapply base_contextual_step; eapply UnOpS; eauto. + * destruct H2 as [e' H2]. eexists. eauto. + - (* pair *) + destruct (IH2 HeqΓ Heqn) as [H2|H2]; [destruct (IH1 HeqΓ Heqn) as [H1|H1]|]. + + left. done. + + right. destruct H1 as [e1' Hstep]. eexists. eauto. + + right. destruct H2 as [e2' H2]. eexists. eauto. + - (* fst *) + destruct (IH HeqΓ Heqn) as [H | H]. + + eapply canonical_values_prod in Hty as (e1 & e2 & -> & ? & ?); last done. + right. eexists. eapply base_contextual_step. econstructor; done. + + right. destruct H as [e' H]. eexists. eauto. + - (* snd *) + destruct (IH HeqΓ Heqn) as [H | H]. + + eapply canonical_values_prod in Hty as (e1 & e2 & -> & ? & ?); last done. + right. eexists. eapply base_contextual_step. econstructor; done. + + right. destruct H as [e' H]. eexists. eauto. + - (* injl *) + destruct (IH HeqΓ Heqn) as [H | H]. + + left. done. + + right. destruct H as [e' H]. eexists. eauto. + - (* injr *) + destruct (IH HeqΓ Heqn) as [H | H]. + + left. done. + + right. destruct H as [e' H]. eexists. eauto. + - (* case *) + right. destruct (IHe HeqΓ Heqn) as [H1|H1]. + + eapply canonical_values_sum in Htye as [(e' & -> & ?) | (e' & -> & ?)]; last done. + * eexists. eapply base_contextual_step. econstructor. done. + * eexists. eapply base_contextual_step. econstructor. done. + + destruct H1 as [e' H1]. eexists. eauto. +Qed. + + + + +Definition ectx_typing (K: ectx) (A B: type) := + ∀ e, TY 0; ∅ ⊢ e : A → TY 0; ∅ ⊢ (fill K e) : B. + + +Lemma fill_typing_decompose K e A: + TY 0; ∅ ⊢ fill K e : A → + ∃ B, TY 0; ∅ ⊢ e : B ∧ ectx_typing K B A. +Proof. + unfold ectx_typing; induction K in A |-*; simpl; inversion 1; subst; eauto. + all: edestruct IHK as (? & ? & ?); eauto. +Qed. + +Lemma fill_typing_compose K e A B: + TY 0; ∅ ⊢ e : B → + ectx_typing K B A → + TY 0; ∅ ⊢ fill K e : A. +Proof. + intros H1 H2; by eapply H2. +Qed. + +Lemma fmap_up_subst σ Γ: ⤉(subst σ <$> Γ) = subst (up σ) <$> ⤉Γ. +Proof. + rewrite -!map_fmap_compose. + eapply map_fmap_ext. intros x A _. by asimpl. +Qed. + +Lemma typed_subst_type n m Γ e A σ: + TY n; Γ ⊢ e : A → (∀ k, k < n → type_wf m (σ k)) → TY m; (subst σ) <$> Γ ⊢ e : A.[σ]. +Proof. + induction 1 as [ n Γ x A Heq | | | n Γ e A Hty IH | |n Γ A B e Hwf Hwf' Hty IH | n Γ A B e e' x Hwf Hty1 IH1 Hty2 IH2 | | | | | |? ? ? ? ? ? ? ? Hop | ? ? ? ? ? ? Hop | | | | | | ] in σ, m |-*; simpl; intros Hlt; eauto. + - econstructor. rewrite lookup_fmap Heq //=. + - econstructor; last by eapply type_wf_subst. + rewrite -fmap_insert. eauto. + - econstructor; last by eapply type_wf_subst. eauto. + - econstructor. rewrite fmap_up_subst. eapply IH. + intros [| x] Hlt'; rewrite /up //=. + + econstructor. lia. + + eapply type_wf_rename; last by (intros ???; simpl; lia). + eapply Hlt. lia. + - replace (A.[B/].[σ]) with (A.[up σ].[B.[σ]/]) by by asimpl. + eauto using type_wf_subst. + - (* pack *) + eapply (typed_pack _ _ _ (subst σ B)). + + eapply type_wf_subst; done. + + eapply type_wf_subst; first done. + intros [ | k] Hk; first ( asimpl;constructor; lia). + rewrite /up //=. eapply type_wf_rename; last by (intros ???; simpl; lia). + eapply Hlt. lia. + + replace (A.[up σ].[B.[σ]/]) with (A.[B/].[σ]) by by asimpl. + eauto using type_wf_subst. + - (* unpack *) + eapply (typed_unpack _ _ A.[up σ]). + + eapply type_wf_subst; done. + + replace (∃: A.[up σ])%ty with ((∃: A).[σ])%ty by by asimpl. + eapply IH1. done. + + rewrite fmap_up_subst. rewrite -fmap_insert. + replace (B.[σ].[ren (+1)]) with (B.[ren(+1)].[up σ]) by by asimpl. + eapply IH2. + intros [ | k] Hk; asimpl; first (constructor; lia). + eapply type_wf_subst; first (eapply Hlt; lia). + intros k' Hk'. asimpl. constructor. lia. + - (* binop *) + inversion Hop; subst. + all: econstructor; naive_solver. + - (* unop *) + inversion Hop; subst. + all: econstructor; naive_solver. + - econstructor; last naive_solver. by eapply type_wf_subst. + - econstructor; last naive_solver. by eapply type_wf_subst. +Qed. + +Lemma typed_subst_type_closed C e A: + type_wf 0 C → TY 1; ⤉∅ ⊢ e : A → TY 0; ∅ ⊢ e : A.[C/]. +Proof. + intros Hwf Hty. eapply typed_subst_type with (σ := C .: ids) (m := 0) in Hty; last first. + { intros [|k] Hlt; last lia. done. } + revert Hty. by rewrite !fmap_empty. +Qed. + +Lemma typed_subst_type_closed' x C B e A: + type_wf 0 A → + type_wf 1 C → + type_wf 0 B → + TY 1; <[x := C]> ∅ ⊢ e : A → + TY 0; <[x := C.[B/]]> ∅ ⊢ e : A. +Proof. + intros ??? Hty. + set (s := (subst (B.:ids))). + rewrite -(fmap_empty s) -(fmap_insert s). + replace A with (A.[B/]). + 2: { replace A with (A.[ids]) at 2 by by asimpl. + eapply type_wf_subst_dom; first done. lia. + } + eapply typed_subst_type; first done. + intros [ | k] Hk; last lia. done. +Qed. + +Lemma typed_preservation_base_step e e' A: + TY 0; ∅ ⊢ e : A → + base_step e e' → + TY 0; ∅ ⊢ e' : A. +Proof. + intros Hty Hstep. destruct Hstep as [ | | | op e v v' Heq Heval | op e1 v1 e2 v2 v3 Heq1 Heq2 Heval | | | | | | ]; subst. + - eapply app_inversion in Hty as (B & H1 & H2). + destruct x as [|x]. + { eapply lam_anon_inversion in H1 as (C & D & [= -> ->] & Hwf & Hty). done. } + eapply lam_inversion in H1 as (C & D & Heq & Hwf & Hty). + injection Heq as -> ->. + eapply typed_substitutivity; eauto. + - eapply type_app_inversion in Hty as (B & C & -> & Hwf & Hty). + eapply type_lam_inversion in Hty as (A & Heq & Hty). + injection Heq as ->. by eapply typed_subst_type_closed. + - (* unpack *) + eapply type_unpack_inversion in Hty as (B & x' & -> & Hwf & Hty1 & Hty2). + eapply type_pack_inversion in Hty1 as (B' & C & [= <-] & Hty1 & ? & ?). + eapply typed_substitutivity. + { apply Hty1. } + rewrite fmap_empty in Hty2. + eapply typed_subst_type_closed'; eauto. + replace A with A.[ids] by by asimpl. + enough (A.[ids] = A.[ren (+1)]) as -> by done. + eapply type_wf_subst_dom; first done. intros; lia. + - (* unop *) + eapply unop_inversion in Hty as (A1 & Hop & Hty). + assert ((A1 = Int ∧ A = Int) ∨ (A1 = Bool ∧ A = Bool)) as [(-> & ->) | (-> & ->)]. + { inversion Hop; subst; eauto. } + + eapply canonical_values_int in Hty as [n ->]; last by eapply is_val_spec; eauto. + simpl in Heq. injection Heq as <-. + inversion Hop; subst; simpl in *; injection Heval as <-; constructor. + + eapply canonical_values_bool in Hty as [b ->]; last by eapply is_val_spec; eauto. + simpl in Heq. injection Heq as <-. + inversion Hop; subst; simpl in *; injection Heval as <-; constructor. + - (* binop *) + eapply binop_inversion in Hty as (A1 & A2 & Hop & Hty1 & Hty2). + assert (A1 = Int ∧ A2 = Int ∧ (A = Int ∨ A = Bool)) as (-> & -> & HC). + { inversion Hop; subst; eauto. } + eapply canonical_values_int in Hty1 as [n ->]; last by eapply is_val_spec; eauto. + eapply canonical_values_int in Hty2 as [m ->]; last by eapply is_val_spec; eauto. + simpl in Heq1, Heq2. injection Heq1 as <-. injection Heq2 as <-. + simpl in Heval. + inversion Hop; subst; simpl in *; injection Heval as <-; constructor. + - by eapply if_inversion in Hty as (H1 & H2 & H3). + - by eapply if_inversion in Hty as (H1 & H2 & H3). + - by eapply fst_inversion in Hty as (B & (? & ? & [= <- <-] & ? & ?)%pair_inversion). + - by eapply snd_inversion in Hty as (B & (? & ? & [= <- <-] & ? & ?)%pair_inversion). + - eapply case_inversion in Hty as (B & C & (? & ? & [= <- <-] & Hty & ?)%injl_inversion & ? & ?). + eauto. + - eapply case_inversion in Hty as (B & C & (? & ? & [= <- <-] & Hty & ?)%injr_inversion & ? & ?). + eauto. +Qed. + +Lemma typed_preservation e e' A: + TY 0; ∅ ⊢ e : A → + contextual_step e e' → + TY 0; ∅ ⊢ e' : A. +Proof. + intros Hty Hstep. destruct Hstep as [K e1 e2 -> -> Hstep]. + eapply fill_typing_decompose in Hty as [B [H1 H2]]. + eapply fill_typing_compose; last done. + by eapply typed_preservation_base_step. +Qed. + +Lemma typed_safety e1 e2 A: + TY 0; ∅ ⊢ e1 : A → + rtc contextual_step e1 e2 → + is_val e2 ∨ reducible e2. +Proof. + induction 2; eauto using typed_progress, typed_preservation. +Qed. + +(** derived typing rules *) +Lemma typed_tapp' n Γ A B C e : + TY n; Γ ⊢ e : (∀: A) → + type_wf n B → + C = A.[B/] → + TY n; Γ ⊢ e <> : C. +Proof. + intros; subst C; by eapply typed_tapp. +Qed. diff --git a/theories/type_systems/systemf_mu/exercises06.v b/theories/type_systems/systemf_mu/exercises06.v new file mode 100644 index 0000000..4bf529e --- /dev/null +++ b/theories/type_systems/systemf_mu/exercises06.v @@ -0,0 +1,202 @@ +From stdpp Require Import gmap base relations tactics. +From iris Require Import prelude. +From semantics.ts.systemf_mu Require Import lang notation types pure tactics logrel. +From Autosubst Require Import Autosubst. + +(** * Exercise Sheet 6 *) + +Notation sub := lang.subst. + +Implicit Types + (e : expr) + (v : val) + (A B : type) +. + +(** ** Exercise 5: Keep Rollin' *) +(** This exercise is slightly tricky. + We strongly recommend you to first work on the other exercises. + You may use the results from this exercise, in particular the fixpoint combinator and its typing, in other exercises, however (which is why it comes first in this Coq file). + *) +Section recursion_combinator. + Variable (f x: string). (* the template of the recursive function *) + Hypothesis (Hfx: f ≠ x). + + (** Recursion Combinator *) + (* First, setup a definition [Rec] satisfying the reduction lemmas below. *) + + (** You may find an auxiliary definition [rec_body] helpful *) + Definition rec_body (t: expr) : expr := + roll (λ: f x, #0). (* TODO *) + + Definition Rec (t: expr): val := + λ: x, rec_body t. (* TODO *) + + Lemma closed_rec_body t : + is_closed [] t → is_closed [] (rec_body t). + Proof. simplify_closed. Qed. + Lemma closed_Rec t : + is_closed [] t → is_closed [] (Rec t). + Proof. simplify_closed. Qed. + Lemma is_val_Rec t: + is_val (Rec t). + Proof. done. Qed. + + + + Lemma Rec_red (t e: expr): + is_val e → + is_val t → + is_closed [] e → + is_closed [] t → + rtc contextual_step ((Rec t) e) (t (Rec t) e). + Proof. + (* TODO: exercise *) + Admitted. + + + Lemma rec_body_typing n Γ (A B: type) t : + Γ !! x = None → + Γ !! f = None → + type_wf n A → + type_wf n B → + TY n; Γ ⊢ t : ((A → B) → (A → B)) → + TY n; Γ ⊢ rec_body t : (μ: #0 → rename (+1) A → rename (+1) B). + Proof. + (* TODO: exercise *) + Admitted. + + + Lemma Rec_typing n Γ A B t: + type_wf n A → + type_wf n B → + Γ !! x = None → + Γ !! f = None → + TY n; Γ ⊢ t : ((A → B) → (A → B)) → + TY n; Γ ⊢ (Rec t) : (A → B). + Proof. + (* TODO: exercise *) + Admitted. + + +End recursion_combinator. + +Definition Fix (f x: string) (e: expr) : val := (Rec f x (Lam f%string (Lam x%string e))). +(** We "seal" the definition to make it opaque to the [solve_typing] tactic. + We do not want [solve_typing] to unfold the definition, instead, we should manually + apply the typing rule below. + + You do not need to understand this in detail -- the only thing of relevance is that + you can use the equality [Fix_eq] to unfold the definition of [Fix]. +*) +Definition Fix_aux : seal (Fix). Proof. by eexists. Qed. +Definition Fix' := Fix_aux.(unseal). +Lemma Fix_eq : Fix' = Fix. +Proof. rewrite -Fix_aux.(seal_eq) //. Qed. +Arguments Fix' _ _ _ : simpl never. + +(* the actual fixpoint combinator *) +Notation "'fix:' f x := e" := (Fix' f x e)%E + (at level 200, f, x at level 1, e at level 200, + format "'[' 'fix:' f x := '/ ' e ']'") : val_scope. +Notation "'fix:' f x := e" := (Fix' f x e)%E + (at level 200, f, x at level 1, e at level 200, + format "'[' 'fix:' f x := '/ ' e ']'") : expr_scope. + + +Lemma fix_red (f x: string) (e e': expr): + is_closed [x; f] e → + is_closed [] e' → + is_val e' → + f ≠ x → + rtc contextual_step ((fix: f x := e) e')%V (sub x e' (sub f (fix: f x := e)%V e)). +Proof. + (* TODO: exercise *) +Admitted. + + +Lemma fix_typing n Γ (f x: string) (A B: type) (e: expr): + type_wf n A → + type_wf n B → + f ≠ x → + TY n; <[x := A]> (<[f := (A → B)%ty]> Γ) ⊢ e : B → + TY n; Γ ⊢ (fix: f x := e) : (A → B). +Proof. + (* TODO: exercise *) +Admitted. + + +(** ** Exercise 1: Encode arithmetic expressions *) + +Definition aexpr : type := #0 (* TODO *). + +Definition num_val (v : val) : val := #0 (* TODO *). +Definition num_expr (e : expr) : expr := #0 (* TODO *). + +Definition plus_val (v1 v2 : val) : val := #0 (* TODO *). +Definition plus_expr (e1 e2 : expr) : expr := #0 (* TODO *). + +Definition mul_val (v1 v2 : val) : val := #0 (* TODO *). +Definition mul_expr (e1 e2 : expr) : expr := #0 (* TODO *). + +Lemma num_expr_typed n Γ e : + TY n; Γ ⊢ e : Int → + TY n; Γ ⊢ num_expr e : aexpr. +Proof. + intros. solve_typing. + (* TODO: exercise *) +Admitted. + + +Lemma plus_expr_typed n Γ e1 e2 : + TY n; Γ ⊢ e1 : aexpr → + TY n; Γ ⊢ e2 : aexpr → + TY n; Γ ⊢ plus_expr e1 e2 : aexpr. +Proof. + (*intros; solve_typing.*) + (* TODO: exercise *) +Admitted. + + +Lemma mul_expr_typed n Γ e1 e2 : + TY n; Γ ⊢ e1 : aexpr → + TY n; Γ ⊢ e2 : aexpr → + TY n; Γ ⊢ mul_expr e1 e2 : aexpr. +Proof. + (*intros; solve_typing.*) + (* TODO: exercise *) +Admitted. + + +Definition eval_aexpr : val := + #0. (* TODO *) + +Lemma eval_aexpr_typed Γ n : + TY n; Γ ⊢ eval_aexpr : (aexpr → Int). +Proof. + (* TODO: exercise *) +Admitted. + + + +(** Exercise 3: Lists *) + +Definition list_t (A : type) : type := + ∃: (#0 (* mynil *) + × (A.[ren (+1)] → #0 → #0) (* mycons *) + × (∀: #1 → #0 → (A.[ren (+2)] → #1 → #0) → #0)) (* mylistcase *) + . + +Definition mylist_impl : val := + #0 (* TODO *) + . + + + +Lemma mylist_impl_sem_typed A : + type_wf 0 A → + ∀ k, 𝒱 (list_t A) δ_any k mylist_impl. +Proof. + (* TODO: exercise *) +Admitted. + diff --git a/theories/type_systems/systemf_mu/lang.v b/theories/type_systems/systemf_mu/lang.v new file mode 100644 index 0000000..6261b43 --- /dev/null +++ b/theories/type_systems/systemf_mu/lang.v @@ -0,0 +1,729 @@ +From stdpp Require Export binders strings. +From iris.prelude Require Import options. +From semantics.lib Require Export maps. + +Declare Scope expr_scope. +Declare Scope val_scope. +Delimit Scope expr_scope with E. +Delimit Scope val_scope with V. + +(** Expressions and vals. *) +Inductive base_lit : Set := + | LitInt (n : Z) | LitBool (b : bool) | LitUnit. +Inductive un_op : Set := + | NegOp | MinusUnOp. +Inductive bin_op : Set := + | PlusOp | MinusOp | MultOp (* Arithmetic *) + | LtOp | LeOp | EqOp. (* Comparison *) + + +Inductive expr := + | Lit (l : base_lit) + (* Base lambda calculus *) + | Var (x : string) + | Lam (x : binder) (e : expr) + | App (e1 e2 : expr) + (* Base types and their operations *) + | UnOp (op : un_op) (e : expr) + | BinOp (op : bin_op) (e1 e2 : expr) + | If (e0 e1 e2 : expr) + (* Polymorphism *) + | TApp (e : expr) + | TLam (e : expr) + | Pack (e : expr) + | Unpack (x : binder) (e1 e2 : expr) + (* Products *) + | Pair (e1 e2 : expr) + | Fst (e : expr) + | Snd (e : expr) + (* Sums *) + | InjL (e : expr) + | InjR (e : expr) + | Case (e0 : expr) (e1 : expr) (e2 : expr) + (* Isorecursive types *) + | Roll (e : expr) + | Unroll (e : expr) +. + +Bind Scope expr_scope with expr. + +Inductive val := + | LitV (l : base_lit) + | LamV (x : binder) (e : expr) + | TLamV (e : expr) + | PackV (v : val) + | PairV (v1 v2 : val) + | InjLV (v : val) + | InjRV (v : val) + | RollV (v : val) +. + +Bind Scope val_scope with val. + +Fixpoint of_val (v : val) : expr := + match v with + | LitV l => Lit l + | LamV x e => Lam x e + | TLamV e => TLam e + | PackV v => Pack (of_val v) + | PairV v1 v2 => Pair (of_val v1) (of_val v2) + | InjLV v => InjL (of_val v) + | InjRV v => InjR (of_val v) + | RollV v => Roll (of_val v) + end. + +Fixpoint to_val (e : expr) : option val := + match e with + | Lit l => Some $ LitV l + | Lam x e => Some (LamV x e) + | TLam e => Some (TLamV e) + | Pack e => + to_val e ≫= (λ v, Some $ PackV v) + | Pair e1 e2 => + to_val e1 ≫= (λ v1, to_val e2 ≫= (λ v2, Some $ PairV v1 v2)) + | InjL e => to_val e ≫= (λ v, Some $ InjLV v) + | InjR e => to_val e ≫= (λ v, Some $ InjRV v) + | Roll e => to_val e ≫= (λ v, Some $ RollV v) + | _ => None + end. + +(** Equality and other typeclass stuff *) +Lemma to_of_val v : to_val (of_val v) = Some v. +Proof. + by induction v; simplify_option_eq; repeat f_equal; try apply (proof_irrel _). +Qed. + +Lemma of_to_val e v : to_val e = Some v → of_val v = e. +Proof. + revert v; induction e; intros v ?; simplify_option_eq; auto with f_equal. +Qed. + +#[export] Instance of_val_inj : Inj (=) (=) of_val. +Proof. by intros ?? Hv; apply (inj Some); rewrite <-!to_of_val, Hv. Qed. + +(** structural computational version *) +Fixpoint is_val (e : expr) : Prop := + match e with + | Lit l => True + | Lam x e => True + | TLam e => True + | Pack e => is_val e + | Pair e1 e2 => is_val e1 ∧ is_val e2 + | InjL e => is_val e + | InjR e => is_val e + | Roll e => is_val e + | _ => False + end. +Lemma is_val_spec e : is_val e ↔ ∃ v, to_val e = Some v. +Proof. + induction e as [ | | ? e IH | e1 IH1 e2 IH2 | e IH | ? e1 IH1 e2 IH2 | e1 IH1 e2 IH2 e3 IH3 | e IH | e IH | e IH | ? e1 IH1 e2 IH2 | e1 IH1 e2 IH2 | e IH | e IH | e IH | e IH | e1 IH1 e2 IH2 e3 IH3 | e IH | e IH]; + simpl; (split; [ | intros (v & Heq)]); simplify_option_eq; try naive_solver. + - rewrite IH. intros (v & ->). eauto. + - rewrite IH1, IH2. intros [(v1 & ->) (v2 & ->)]. eauto. + - rewrite IH. intros (v & ->). eauto. + - rewrite IH. intros (v & ->); eauto. + - rewrite IH. intros (v & ->). eauto. +Qed. + +Global Instance base_lit_eq_dec : EqDecision base_lit. +Proof. solve_decision. Defined. +Global Instance un_op_eq_dec : EqDecision un_op. +Proof. solve_decision. Defined. +Global Instance bin_op_eq_dec : EqDecision bin_op. +Proof. solve_decision. Defined. + + +(** Substitution *) +Fixpoint subst (x : string) (es : expr) (e : expr) : expr := + match e with + | Lit _ => e + | Var y => if decide (x = y) then es else Var y + | Lam y e => + Lam 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) + | TApp e => TApp (subst x es e) + | TLam e => TLam (subst x es e) + | Pack e => Pack (subst x es e) + | Unpack y e1 e2 => Unpack y (subst x es e1) (if decide (BNamed x = y) then e2 else subst x es e2) + | UnOp op e => UnOp op (subst x es e) + | BinOp op e1 e2 => BinOp op (subst x es e1) (subst x es e2) + | If e0 e1 e2 => If (subst x es e0) (subst x es e1) (subst x es e2) + | Pair e1 e2 => Pair (subst x es e1) (subst x es e2) + | Fst e => Fst (subst x es e) + | Snd e => Snd (subst x es e) + | InjL e => InjL (subst x es e) + | InjR e => InjR (subst x es e) + | Case e0 e1 e2 => Case (subst x es e0) (subst x es e1) (subst x es e2) + | Roll e => Roll (subst x es e) + | Unroll e => Unroll (subst x es e) + end. + +Definition subst' (mx : binder) (es : expr) : expr → expr := + match mx with BNamed x => subst x es | BAnon => id end. + +(** The stepping relation *) +Definition un_op_eval (op : un_op) (v : val) : option val := + match op, v with + | NegOp, LitV (LitBool b) => Some $ LitV $ LitBool (negb b) + | MinusUnOp, LitV (LitInt n) => Some $ LitV $ LitInt (- n) + | _, _ => None + end. + +Definition bin_op_eval_int (op : bin_op) (n1 n2 : Z) : option base_lit := + match op with + | PlusOp => Some $ LitInt (n1 + n2) + | MinusOp => Some $ LitInt (n1 - n2) + | MultOp => Some $ LitInt (n1 * n2) + | LtOp => Some $ LitBool (bool_decide (n1 < n2)) + | LeOp => Some $ LitBool (bool_decide (n1 ≤ n2)) + | EqOp => Some $ LitBool (bool_decide (n1 = n2)) + end%Z. + +Definition bin_op_eval (op : bin_op) (v1 v2 : val) : option val := + match v1, v2 with + | LitV (LitInt n1), LitV (LitInt n2) => LitV <$> bin_op_eval_int op n1 n2 + | _, _ => None + end. + +Inductive base_step : expr → expr → Prop := + | BetaS x e1 e2 e' : + is_val e2 → + e' = subst' x e2 e1 → + base_step (App (Lam x e1) e2) e' + | TBetaS e1 : + base_step (TApp (TLam e1)) e1 + | UnpackS e1 e2 e' x : + is_val e1 → + e' = subst' x e1 e2 → + base_step (Unpack x (Pack e1) e2) e' + | UnOpS op e v v' : + to_val e = Some v → + un_op_eval op v = Some v' → + base_step (UnOp op e) (of_val v') + | BinOpS op e1 v1 e2 v2 v' : + to_val e1 = Some v1 → + to_val e2 = Some v2 → + bin_op_eval op v1 v2 = Some v' → + base_step (BinOp op e1 e2) (of_val v') + | IfTrueS e1 e2 : + base_step (If (Lit $ LitBool true) e1 e2) e1 + | IfFalseS e1 e2 : + base_step (If (Lit $ LitBool false) e1 e2) e2 + | FstS e1 e2 : + is_val e1 → + is_val e2 → + base_step (Fst (Pair e1 e2)) e1 + | SndS e1 e2 : + is_val e1 → + is_val e2 → + base_step (Snd (Pair e1 e2)) e2 + | CaseLS e e1 e2 : + is_val e → + base_step (Case (InjL e) e1 e2) (App e1 e) + | CaseRS e e1 e2 : + is_val e → + base_step (Case (InjR e) e1 e2) (App e2 e) + | UnrollS e : + is_val e → + base_step (Unroll (Roll e)) e +. + +(* Misc *) +Lemma is_val_of_val v : is_val (of_val v). +Proof. apply is_val_spec. rewrite to_of_val. eauto. Qed. + +(** If [e1] makes a base step to a value under some state [σ1] then any base + step from [e1] under any other state [σ1'] must necessarily be to a value. *) +Lemma base_step_to_val e1 e2 e2' : + base_step e1 e2 → + base_step e1 e2' → is_Some (to_val e2) → is_Some (to_val e2'). +Proof. destruct 1; inversion 1; naive_solver. Qed. + +(** We define evaluation contexts *) +(** In contrast to before, we formalize contexts differently in a non-recursive way. + Evaluation contexts are now lists of evaluation context items, and we specify how to + plug these items together when filling the context. + For instance: + - HoleCtx becomes [], the empty list of context items + - e1 ● becomes [AppRCtx e1] + - e1 (● + v2) becomes [PlusLCtx v2; AppRCtx e1] + *) +Inductive ectx_item := + | AppLCtx (v2 : val) + | AppRCtx (e1 : expr) + | TAppCtx + | PackCtx + | UnpackCtx (x : binder) (e2 : expr) + | UnOpCtx (op : un_op) + | BinOpLCtx (op : bin_op) (v2 : val) + | BinOpRCtx (op : bin_op) (e1 : expr) + | IfCtx (e1 e2 : expr) + | PairLCtx (v2 : val) + | PairRCtx (e1 : expr) + | FstCtx + | SndCtx + | InjLCtx + | InjRCtx + | CaseCtx (e1 : expr) (e2 : expr) + | UnrollCtx + | RollCtx +. + +Definition fill_item (Ki : ectx_item) (e : expr) : expr := + match Ki with + | AppLCtx v2 => App e (of_val v2) + | AppRCtx e1 => App e1 e + | TAppCtx => TApp e + | PackCtx => Pack e + | UnpackCtx x e2 => Unpack x e e2 + | UnOpCtx op => UnOp op e + | BinOpLCtx op v2 => BinOp op e (of_val v2) + | BinOpRCtx op e1 => BinOp op e1 e + | IfCtx e1 e2 => If e e1 e2 + | PairLCtx v2 => Pair e (of_val v2) + | PairRCtx e1 => Pair e1 e + | FstCtx => Fst e + | SndCtx => Snd e + | InjLCtx => InjL e + | InjRCtx => InjR e + | CaseCtx e1 e2 => Case e e1 e2 + | UnrollCtx => Unroll e + | RollCtx => Roll e + end. +(* The main [fill] operation is defined below. *) + +(** Basic properties about the language *) +Global Instance fill_item_inj Ki : Inj (=) (=) (fill_item Ki). +Proof. induction Ki; intros ???; simplify_eq/=; auto with f_equal. Qed. + +Lemma fill_item_val Ki e : + is_Some (to_val (fill_item Ki e)) → is_Some (to_val e). +Proof. intros [v ?]. induction Ki; simplify_option_eq; eauto. Qed. + +Lemma val_base_stuck e1 e2 : base_step e1 e2 → to_val e1 = None. +Proof. destruct 1; naive_solver. Qed. + +Lemma of_val_lit v l : + of_val v = (Lit l) → v = LitV l. +Proof. destruct v; simpl; inversion 1; done. Qed. +Lemma of_val_pair_inv (v v1 v2 : val) : + of_val v = Pair (of_val v1) (of_val v2) → v = PairV v1 v2. +Proof. + destruct v; simpl; inversion 1. + apply of_val_inj in H1. apply of_val_inj in H2. subst; done. +Qed. +Lemma of_val_injl_inv (v v' : val) : + of_val v = InjL (of_val v') → v = InjLV v'. +Proof. + destruct v; simpl; inversion 1. + apply of_val_inj in H1. subst; done. +Qed. +Lemma of_val_injr_inv (v v' : val) : + of_val v = InjR (of_val v') → v = InjRV v'. +Proof. + destruct v; simpl; inversion 1. + apply of_val_inj in H1. subst; done. +Qed. +Ltac simplify_val := + repeat match goal with + | H: to_val (of_val ?v) = ?o |- _ => rewrite to_of_val in H + | H: is_val ?e |- _ => destruct (proj1 (is_val_spec e) H) as (? & ?); clear H + | H: to_val ?e = ?v |- _ => is_var e; specialize (of_to_val _ _ H); intros <-; clear H + | H: of_val ?v = Lit ?l |- _ => is_var v; specialize (of_val_lit _ _ H); intros ->; clear H + | |- context[to_val (of_val ?e)] => rewrite to_of_val + end. + +Lemma base_ectxi_step_val Ki e e2 : + base_step (fill_item Ki e) e2 → is_Some (to_val e). +Proof. + revert e2. induction Ki; inversion_clear 1; simplify_option_eq; simplify_val; eauto. +Qed. + +Lemma fill_item_no_val_inj Ki1 Ki2 e1 e2 : + to_val e1 = None → to_val e2 = None → + fill_item Ki1 e1 = fill_item Ki2 e2 → Ki1 = Ki2. +Proof. + revert Ki1. induction Ki2; intros Ki1; induction Ki1; try naive_solver eauto with f_equal. + all: intros ?? Heq; injection Heq; intros ??; simplify_eq; simplify_val; naive_solver. +Qed. + +Section contexts. + Notation ectx := (list ectx_item). + Definition empty_ectx : ectx := []. + Definition comp_ectx : ectx → ectx → ectx := flip (++). + + Definition fill (K : ectx) (e : expr) : expr := foldl (flip fill_item) e K. + Lemma fill_app (K1 K2 : ectx) e : fill (K1 ++ K2) e = fill K2 (fill K1 e). + Proof. apply foldl_app. Qed. + Lemma fill_empty e : fill empty_ectx e = e. + Proof. done. Qed. + Lemma fill_comp K1 K2 e : fill K1 (fill K2 e) = fill (comp_ectx K1 K2) e. + Proof. unfold fill, comp_ectx; simpl. by rewrite foldl_app. Qed. + Global Instance fill_inj K : Inj (=) (=) (fill K). + Proof. induction K as [|Ki K IH]; unfold Inj; naive_solver. Qed. + + Lemma fill_val K e : is_Some (to_val (fill K e)) → is_Some (to_val e). + Proof. + induction K as [|Ki K IH] in e |-*; [ done |]. by intros ?%IH%fill_item_val. + Qed. + Lemma fill_not_val K e : to_val e = None → to_val (fill K e) = None. + Proof. rewrite !eq_None_not_Some. eauto using fill_val. Qed. + +End contexts. + +(** Contextual steps *) +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 base_contextual_step e1 e2 : + base_step e1 e2 → contextual_step e1 e2. +Proof. apply Ectx_step with empty_ectx; by rewrite ?fill_empty. Qed. + +Lemma fill_contextual_step K e1 e2 : + contextual_step e1 e2 → contextual_step (fill K e1) (fill K e2). +Proof. + destruct 1 as [K' e1' e2' -> ->]. + rewrite !fill_comp. by econstructor. +Qed. + +(** *** closedness *) +Fixpoint is_closed (X : list string) (e : expr) : bool := + match e with + | Var x => bool_decide (x ∈ X) + | Lam x e => is_closed (x :b: X) e + | Unpack x e1 e2 => is_closed X e1 && is_closed (x :b: X) e2 + | Lit _ => true + | UnOp _ e | Fst e | Snd e | InjL e | InjR e | TApp e | TLam e | Pack e | Roll e | Unroll e => is_closed X e + | App e1 e2 | BinOp _ e1 e2 | Pair e1 e2 => is_closed X e1 && is_closed X e2 + | If e0 e1 e2 | Case e0 e1 e2 => + is_closed X e0 && is_closed X e1 && is_closed X e2 + end. + +(** [closed] states closedness as a Coq proposition, through the [Is_true] transformer. *) +Definition closed (X : list string) (e : expr) : Prop := Is_true (is_closed X e). +#[export] Instance closed_proof_irrel X e : ProofIrrel (closed X e). +Proof. unfold closed. apply _. Qed. +#[export] Instance closed_dec X e : Decision (closed X e). +Proof. unfold closed. apply _. Defined. + +(** closed expressions *) +Lemma is_closed_weaken X Y e : is_closed X e → X ⊆ Y → is_closed Y e. +Proof. revert X Y; induction e; naive_solver (eauto; set_solver). Qed. + +Lemma is_closed_weaken_nil X e : is_closed [] e → is_closed X e. +Proof. intros. by apply is_closed_weaken with [], list_subseteq_nil. Qed. + +Lemma is_closed_subst X e x es : + is_closed [] es → is_closed (x :: X) e → is_closed X (subst x es e). +Proof. + intros ?. + induction e in X |-*; simpl; intros ?; destruct_and?; split_and?; simplify_option_eq; + try match goal with + | H : ¬(_ ∧ _) |- _ => apply not_and_l in H as [?%dec_stable|?%dec_stable] + end; eauto using is_closed_weaken with set_solver. +Qed. +Lemma is_closed_do_subst' X e x es : + is_closed [] es → is_closed (x :b: X) e → is_closed X (subst' x es e). +Proof. destruct x; eauto using is_closed_subst. Qed. + +Lemma closed_is_closed X e : + is_closed X e = true ↔ closed X e. +Proof. + unfold closed. split. + - apply Is_true_eq_left. + - apply Is_true_eq_true. +Qed. + +(** Substitution lemmas *) +Lemma subst_is_closed X e x es : is_closed X e → x ∉ X → subst x es e = e. +Proof. + induction e in X |-*; simpl; rewrite ?bool_decide_spec, ?andb_True; intros ??; + repeat case_decide; simplify_eq; simpl; f_equal; intuition eauto with set_solver. +Qed. + +Lemma subst_is_closed_nil e x es : is_closed [] e → subst x es e = e. +Proof. intros. apply subst_is_closed with []; set_solver. Qed. +Lemma subst'_is_closed_nil e x es : is_closed [] e → subst' x es e = e. +Proof. intros. destruct x as [ | x]. { done. } by apply subst_is_closed_nil. Qed. + +(** ** More on the contextual semantics *) + +Definition base_reducible (e : expr) := + ∃ e', base_step e e'. +Definition base_irreducible (e : expr) := + ∀ e', ¬base_step e e'. +Definition base_stuck (e : expr) := + to_val e = None ∧ base_irreducible e. + +(** Given a base redex [e1_redex] somewhere in a term, and another + decomposition of the same term into [fill K' e1'] such that [e1'] is not + a value, then the base redex context is [e1']'s context [K'] filled with + another context [K'']. In particular, this implies [e1 = fill K'' + e1_redex] by [fill_inj], i.e., [e1]' contains the base redex.) + *) +Lemma step_by_val K' K_redex e1' e1_redex e2 : + fill K' e1' = fill K_redex e1_redex → + to_val e1' = None → + base_step e1_redex e2 → + ∃ K'', K_redex = comp_ectx K' K''. +Proof. + rename K' into K. rename K_redex into K'. + rename e1' into e1. rename e1_redex into e1'. + intros Hfill Hred Hstep; revert K' Hfill. + induction K as [|Ki K IH] using rev_ind; simpl; intros K' Hfill; eauto using app_nil_r. + destruct K' as [|Ki' K' _] using @rev_ind; simplify_eq/=. + { rewrite fill_app in Hstep. apply base_ectxi_step_val in Hstep. + apply fill_val in Hstep. by apply not_eq_None_Some in Hstep. } + rewrite !fill_app in Hfill; simpl in Hfill. + assert (Ki = Ki') as ->. + { eapply fill_item_no_val_inj, Hfill. + - by apply fill_not_val. + - apply fill_not_val. eauto using val_base_stuck. + } + simplify_eq. destruct (IH K') as [K'' ->]; auto. + exists K''. unfold comp_ectx; simpl. rewrite assoc; [done |]. apply _. +Qed. + +(** If [fill K e] takes a base step, then either [e] is a value or [K] is + the empty evaluation context. In other words, if [e] is not a value + wrapping it in a context does not add new base redex positions. *) +Lemma base_ectx_step_val K e e2 : + base_step (fill K e) e2 → is_Some (to_val e) ∨ K = empty_ectx. +Proof. + destruct K as [|Ki K _] using rev_ind; simpl; [by auto|]. + rewrite fill_app; simpl. + intros ?%base_ectxi_step_val; eauto using fill_val. +Qed. + +(** If a [contextual_step] preserving a surrounding context [K] happens, + the reduction happens entirely below the context. *) +Lemma contextual_step_ectx_inv K e e' : + to_val e = None → + contextual_step (fill K e) (fill K e') → + contextual_step e e'. +Proof. + intros ? Hcontextual. + inversion Hcontextual; subst. + eapply step_by_val in H as (K'' & Heq); [ | done | done]. + subst K0. + rewrite <-fill_comp in H1. rewrite <-fill_comp in H0. + apply fill_inj in H1. apply fill_inj in H0. subst. + econstructor; done. +Qed. + +Lemma base_reducible_contextual_step_ectx K e1 e2 : + base_reducible e1 → + contextual_step (fill K e1) e2 → + ∃ e2', e2 = fill K e2' ∧ base_step e1 e2'. +Proof. + intros (e2''&HhstepK) [K' e1' e2' HKe1 -> Hstep]. + edestruct (step_by_val K) as [K'' ?]; + eauto using val_base_stuck; simplify_eq/=. + rewrite <-fill_comp in HKe1; simplify_eq. + exists (fill K'' e2'). rewrite fill_comp; split; [done | ]. + apply base_ectx_step_val in HhstepK as [[v ?]|?]; simplify_eq. + { apply val_base_stuck in Hstep; simplify_eq. } + by rewrite !fill_empty. +Qed. + +Lemma base_reducible_contextual_step e1 e2 : + base_reducible e1 → + contextual_step e1 e2 → + base_step e1 e2. +Proof. + intros. + edestruct (base_reducible_contextual_step_ectx empty_ectx) as (?&?&?); + rewrite ?fill_empty; eauto. + by simplify_eq; rewrite fill_empty. +Qed. + +(** *** Reducibility *) +Definition reducible (e : expr) := + ∃ e', contextual_step e e'. +Definition irreducible (e : expr) := + ∀ e', ¬contextual_step e e'. +Definition stuck (e : expr) := + to_val e = None ∧ irreducible e. +Definition not_stuck (e : expr) := + is_Some (to_val e) ∨ reducible e. + +Lemma base_step_not_stuck e e' : base_step e e' → not_stuck e. +Proof. unfold not_stuck, reducible; simpl. eauto 10 using base_contextual_step. Qed. + +Lemma val_stuck e e' : contextual_step e e' → to_val e = None. +Proof. + intros [??? -> -> ?%val_base_stuck]. + apply eq_None_not_Some. by intros ?%fill_val%eq_None_not_Some. +Qed. +Lemma not_reducible e : ¬reducible e ↔ irreducible e. +Proof. unfold reducible, irreducible. naive_solver. Qed. +Lemma reducible_not_val e : reducible e → to_val e = None. +Proof. intros (?&?). eauto using val_stuck. Qed. +Lemma val_irreducible e : is_Some (to_val e) → irreducible e. +Proof. intros [??] ? ?%val_stuck. by destruct (to_val e). Qed. + +Lemma irreducible_fill K e : + irreducible (fill K e) → irreducible e. +Proof. + intros Hirred e' Hstep. + eapply Hirred. by apply fill_contextual_step. +Qed. + +Lemma base_reducible_reducible e : + base_reducible e → reducible e. +Proof. intros (e' & Hhead). exists e'. by apply base_contextual_step. Qed. + + +(* we derive a few lemmas about contextual steps *) +Lemma contextual_step_app_l e1 e1' e2: + is_val e2 → + contextual_step e1 e1' → + contextual_step (App e1 e2) (App e1' e2). +Proof. + intros [v <-%of_to_val]%is_val_spec Hcontextual. + by eapply (fill_contextual_step [AppLCtx _]). +Qed. + +Lemma contextual_step_app_r e1 e2 e2': + contextual_step e2 e2' → + contextual_step (App e1 e2) (App e1 e2'). +Proof. + intros Hcontextual. + by eapply (fill_contextual_step [AppRCtx e1]). +Qed. + +Lemma contextual_step_tapp e e': + contextual_step e e' → + contextual_step (TApp e) (TApp e'). +Proof. + intros Hcontextual. + by eapply (fill_contextual_step [TAppCtx]). +Qed. + +Lemma contextual_step_pack e e': + contextual_step e e' → + contextual_step (Pack e) (Pack e'). +Proof. + intros Hcontextual. + by eapply (fill_contextual_step [PackCtx]). +Qed. + +Lemma contextual_step_unpack x e e' e2: + contextual_step e e' → + contextual_step (Unpack x e e2) (Unpack x e' e2). +Proof. + intros Hcontextual. + by eapply (fill_contextual_step [UnpackCtx x e2]). +Qed. + +Lemma contextual_step_unop op e e': + contextual_step e e' → + contextual_step (UnOp op e) (UnOp op e'). +Proof. + intros Hcontextual. + by eapply (fill_contextual_step [UnOpCtx op]). +Qed. + +Lemma contextual_step_binop_l op e1 e1' e2: + is_val e2 → + contextual_step e1 e1' → + contextual_step (BinOp op e1 e2) (BinOp op e1' e2). +Proof. + intros [v <-%of_to_val]%is_val_spec Hcontextual. + by eapply (fill_contextual_step [BinOpLCtx op v]). +Qed. + +Lemma contextual_step_binop_r op e1 e2 e2': + contextual_step e2 e2' → + contextual_step (BinOp op e1 e2) (BinOp op e1 e2'). +Proof. + intros Hcontextual. + by eapply (fill_contextual_step [BinOpRCtx op e1]). +Qed. + +Lemma contextual_step_if e e' e1 e2: + contextual_step e e' → + contextual_step (If e e1 e2) (If e' e1 e2). +Proof. + intros Hcontextual. + by eapply (fill_contextual_step [IfCtx e1 e2]). +Qed. + +Lemma contextual_step_pair_l e1 e1' e2: + is_val e2 → + contextual_step e1 e1' → + contextual_step (Pair e1 e2) (Pair e1' e2). +Proof. + intros [v <-%of_to_val]%is_val_spec Hcontextual. + by eapply (fill_contextual_step [PairLCtx v]). +Qed. + +Lemma contextual_step_pair_r e1 e2 e2': + contextual_step e2 e2' → + contextual_step (Pair e1 e2) (Pair e1 e2'). +Proof. + intros Hcontextual. + by eapply (fill_contextual_step [PairRCtx e1]). +Qed. + + +Lemma contextual_step_fst e e': + contextual_step e e' → + contextual_step (Fst e) (Fst e'). +Proof. + intros Hcontextual. + by eapply (fill_contextual_step [FstCtx]). +Qed. + +Lemma contextual_step_snd e e': + contextual_step e e' → + contextual_step (Snd e) (Snd e'). +Proof. + intros Hcontextual. + by eapply (fill_contextual_step [SndCtx]). +Qed. + +Lemma contextual_step_injl e e': + contextual_step e e' → + contextual_step (InjL e) (InjL e'). +Proof. + intros Hcontextual. + by eapply (fill_contextual_step [InjLCtx]). +Qed. + +Lemma contextual_step_injr e e': + contextual_step e e' → + contextual_step (InjR e) (InjR e'). +Proof. + intros Hcontextual. + by eapply (fill_contextual_step [InjRCtx]). +Qed. + +Lemma contextual_step_case e e' e1 e2: + contextual_step e e' → + contextual_step (Case e e1 e2) (Case e' e1 e2). +Proof. + intros Hcontextual. + by eapply (fill_contextual_step [CaseCtx e1 e2]). +Qed. + +Lemma contextual_step_roll e e': + contextual_step e e' → + contextual_step (Roll e) (Roll e'). +Proof. by apply (fill_contextual_step [RollCtx]). Qed. + +Lemma contextual_step_unroll e e': + contextual_step e e' → contextual_step (Unroll e) (Unroll e'). +Proof. by apply (fill_contextual_step [UnrollCtx]). Qed. + + +#[export] +Hint Resolve + contextual_step_app_l contextual_step_app_r contextual_step_binop_l contextual_step_binop_r + contextual_step_case contextual_step_fst contextual_step_if contextual_step_injl contextual_step_injr + contextual_step_pack contextual_step_pair_l contextual_step_pair_r contextual_step_snd contextual_step_tapp + contextual_step_tapp contextual_step_unop contextual_step_unpack contextual_step_roll contextual_step_unroll : core. + diff --git a/theories/type_systems/systemf_mu/logrel.v b/theories/type_systems/systemf_mu/logrel.v new file mode 100644 index 0000000..5e4e7ac --- /dev/null +++ b/theories/type_systems/systemf_mu/logrel.v @@ -0,0 +1,1097 @@ +From stdpp Require Import gmap base relations. +From iris Require Import prelude. +From semantics.lib Require Export facts. +From semantics.ts.systemf_mu Require Import lang notation types tactics pure parallel_subst. +From Equations Require Export Equations. + +(** * Logical relation for SystemF + recursive types *) + +Implicit Types + (Δ : nat) + (Γ : typing_context) + (v : val) + (α : var) + (e : expr) + (A : type). + + +(** ** Definition of the logrel *) +(** A semantic type consists of a value-predicate parameterized over a step-index, + a proof of closedness, and a proof of downwards-closure wrt step-indices. *) +Record sem_type := mk_ST { + sem_type_car :> nat → val → Prop; + sem_type_closed_val k v : sem_type_car k v → is_closed [] v; + sem_type_mono : ∀ k k' v, sem_type_car k v → k' ≤ k → sem_type_car k' v + }. + +(** Two tactics we will use later on. + [pose_sem_type P as N] defines a semantic type at name [N] with the value predicate [P]. + *) +(* slightly complicated formulation to make the proof term [p] opaque and prevent it from polluting the context *) +Tactic Notation "pose_sem_type" uconstr(P) "as" ident(N) := + let p := fresh "__p" in + let p2 := fresh "__p2" in + unshelve refine ((λ p p2, let N := (mk_ST P p p2) in _) _ _); first (simpl in p, p2); cycle 1. +Tactic Notation "specialize_sem_type" constr(S) "with" uconstr(P) "as" ident(N) := + pose_sem_type P as N; last specialize (S N). + +(** We represent type variable assignments [δ] as functions [f] into semantic types. + The variable [#n] (in De Bruijn representation) is mapped to [f n]. + *) +Definition tyvar_interp := var → sem_type. +Implicit Types + (δ : tyvar_interp) + (τ : sem_type) +. + +(** + In Coq, we need to make argument why the logical relation is well-defined precise: + (for Coq, that means: we need to show that the recursion is terminating). + + Adding in the recursion over step-indices makes the termination argument slightly more complicated: + we now have a mutual recursion over types, step-indices, and the case of the logrel (the expression relation is defined in terms of the value relation). + + To make this formal, we define a well-founded relation that allows to either decrease the step-index, the type, or switch from the expression case to the value case for recursive calls. + We define size measures for for all three of these things, and then combine them into a big lexicographic ordering [term_rel]. + *) +Equations type_size (A : type) : nat := + type_size Int := 1; + type_size Bool := 1; + type_size Unit := 1; + type_size (A → B) := type_size A + type_size B + 1; + type_size (#α) := 1; + type_size (∀: A) := type_size A + 2; + type_size (∃: A) := type_size A + 2; + type_size (A × B) := type_size A + type_size B + 1; + type_size (A + B) := max (type_size A) (type_size B) + 1; + type_size (μ: A) := type_size A + 2 +. +(* [ltof A R] defines a well-founded measure on type [A] by using a mapping [R] from [A] to [nat] + (it lifts the < relation on natural numbers to [A]) *) +Definition type_lt := ltof type type_size. +#[local] Instance type_lt_wf : WellFounded type_lt. +Proof. apply well_founded_ltof. Qed. + +Inductive val_or_expr : Type := +| inj_val : val → val_or_expr +| inj_expr : expr → val_or_expr. + +Definition val_or_expr_size (ve : val_or_expr) : nat := + match ve with | inj_val _ => 0 | inj_expr _ => 1 end. +Definition val_or_expr_lt := ltof val_or_expr val_or_expr_size. +#[local] Instance val_or_expr_lt_wf : WellFounded val_or_expr_lt. +Proof. apply well_founded_ltof. Qed. + +Definition term_rel := Subterm.lexprod nat (type * val_or_expr) lt (Subterm.lexprod type val_or_expr type_lt val_or_expr_lt). +#[local] Instance term_rel_wf : WellFounded term_rel. apply _. Qed. + +(** *** The logical relation *) +(** Since the relation is step-indexed now, and the argument that the case for recursive types is well-formed + fundamentally requires decreasing the step-index, we also need to convince Equations that this definition is well-formed! + We do this by providing a well-founded termination relation [term_rel] that decreases for each recursive call. + *) +Equations type_interp (ve : val_or_expr) (t : type) δ (k : nat) : Prop + by wf (k, (t, ve)) term_rel := { + + type_interp (inj_val v) Int δ k => + ∃ z : Z, v = #z ; + type_interp (inj_val v) Bool δ k => + ∃ b : bool, v = #b ; + type_interp (inj_val v) Unit δ k => + v = #LitUnit ; + type_interp (inj_val v) (A × B) δ k => + ∃ v1 v2 : val, v = (v1, v2)%V ∧ type_interp (inj_val v1) A δ k ∧ type_interp (inj_val v2) B δ k; + type_interp (inj_val v) (A + B) δ k => + (∃ v' : val, v = InjLV v' ∧ type_interp (inj_val v') A δ k) ∨ + (∃ v' : val, v = InjRV v' ∧ type_interp (inj_val v') B δ k); + + type_interp (inj_val v) (A → B) δ k => + ∃ x e, v = LamV x e ∧ is_closed (x :b: nil) e ∧ + (* We write ∀ (H:k' ≤ k), .. instead of k' ≤ k → .. due to a longstanding Coq quirk, see + https://coq.zulipchat.com/#narrow/stream/237977-Coq-users/topic/.60Program.60.20and.20variable.20names/near/404824378 *) + ∀ v' k' (H:k' ≤ k), + type_interp (inj_val v') A δ k' → + type_interp (inj_expr (subst' x (of_val v') e)) B δ k' ; + type_interp (inj_val v) (#α) δ k => + (δ α).(sem_type_car) k v; + type_interp (inj_val v) (∀: A) δ k => + ∃ e, v = TLamV e ∧ is_closed [] e ∧ + ∀ τ, type_interp (inj_expr e) A (τ .: δ) k; + type_interp (inj_val v) (∃: A) δ k => + ∃ v', v = PackV v' ∧ + ∃ τ : sem_type, type_interp (inj_val v') A (τ .: δ) k; + (** Recursive type case *) + type_interp (inj_val v) (μ: A) δ k => + ∃ v', v = (roll v')%V ∧ is_closed [] v' ∧ ∀ k' (H:k' < k), type_interp (inj_val v') (A.[μ: A/]%ty) δ k'; + + type_interp (inj_expr e) t δ k => + ∀ e' n, n < k → red_nsteps n e e' → ∃ v, to_val e' = Some v ∧ type_interp (inj_val v) t δ (k - n) +}. + +(** Proving that the arguments are decreasing for recursive calls is a bit more messy now, but it's mostly systematic. + Therefore we provide a simple automation tactic that will also become useful a few times below. +*) +Ltac dsimpl := + repeat match goal with + | |- term_rel (?k, _) (?k, _) => + (* step-index is not decreasing, go right *) + right + | |- term_rel (?k1, _) (?k2, _) => + (* use [lia] to decide where to go *) + destruct (decide (k1 < k2)) as [ ? | ?]; [left; lia | assert (k1 = k2) as -> by lia; right] + | |- Subterm.lexprod type val_or_expr _ _ (?t, _) (?t, _) => + (* type is not decreasing, go right *) + right + | |- Subterm.lexprod type val_or_expr _ _ (_, ?a) (_, ?a) => + (* type case is not decreasing, go left *) + left + | |- term_rel (_, _) (_, _) => + (* branch non-deterministically and try to solve the remaining goal *) + first [left; solve [dsimpl] | right; solve [dsimpl]] + | |- Subterm.lexprod type val_or_expr _ _ _ _ => + (* branch non-deterministically *) + first [left; solve [dsimpl] | right; solve [dsimpl]] + | _ => + (* try to solve a leaf, i.e. a [type_lt], [val_or_expr_lt] or [lt] goal *) + unfold val_or_expr_lt, type_lt, ltof; simp type_size; simpl; try lia + end. +(** The tactic solves all of Equations' obligations for showing that the argument decreases. *) +Solve Obligations with (intros; dsimpl). + +(** *** Value relation and expression relation *) +Notation sem_val_rel A δ k v := (type_interp (inj_val v) A δ k). +Notation sem_expr_rel A δ k e := (type_interp (inj_expr e) A δ k). + +Notation 𝒱 A δ k v := (sem_val_rel A δ k v). +Notation ℰ A δ k v := (sem_expr_rel A δ k v). + +Lemma val_rel_is_closed v δ k A: + 𝒱 A δ k v → is_closed [] (of_val v). +Proof. + induction A as [ | | | | | A IHA | | A IH1 B IH2 | A IH1 B IH2 | A IHA] in k, v, δ |-*; simp type_interp. + - by eapply sem_type_closed_val. + - intros [z ->]. done. + - intros [b ->]. done. + - intros ->. done. + - intros (e & -> & ? & _). done. + - intros (v' & -> & (τ & Hinterp)). simpl. by eapply IHA. + - intros (x & e & -> & ? & _). done. + - intros (v1 & v2 & -> & ? & ?). simpl; apply andb_True; split; eauto. + - intros [(v' & -> & ?) | (v' & -> & ?)]; simpl; eauto. + - intros (v' & -> & ? & Ha); done. +Qed. + + +(** This is the Value Inclusion lemma from the lecture notes *) +Lemma sem_val_expr_rel A δ k v : + 𝒱 A δ k v → ℰ A δ k (of_val v). +Proof. + simp type_interp. intros Hv e' n Hn [-> ->]%nsteps_val_inv. + rewrite to_of_val. eexists; split; first done. + replace (k - 0) with k by lia. done. +Qed. +Lemma sem_val_expr_rel' A δ k v e: + to_val e = Some v → + 𝒱 A δ k v → ℰ A δ k e. +Proof. + intros <-%of_to_val. apply sem_val_expr_rel. +Qed. + +Lemma sem_expr_rel_zero_trivial A δ e : + ℰ A δ 0 e. +Proof. + simp type_interp. intros ???. lia. +Qed. + +Lemma sem_expr_rel_of_val A δ k v : + k > 0 → ℰ A δ k (of_val v) → 𝒱 A δ k v. +Proof. + simp type_interp. + intros Hk He. destruct (He (of_val v) 0 ltac:(lia)) as (v' & Hv & Hvr). + { split; first constructor. apply val_irreducible. rewrite to_of_val. eauto. } + rewrite to_of_val in Hv. injection Hv as [= <-]. + replace k with (k - 0) by lia. done. +Qed. + + +(** *** Downwards closure wrt step-index *) +(** Formally proving that the expression and value relations are downwards-closed wrt the step-index + (see the lemmas [val_rel_mono] and [expr_rel_mono] below) is slightly involved: + Intuitively, we'd like to do an inductive argument, since the "base cases" (Int, Bool, etc.) of the + relation are trivially downwards-closed and the other cases inductively preserve this fact. + However, since the relation is defined via recursion on both types and the step-index, we need to + do an induction on both simultaneously. + + For that, we do a well-founded induction with the termination relation [term_rel_wf] we gave to prove + well-formedness of the logical relation. + We can use the inductive hypothesis whenever either the type or the step-index decreases, or we switch from + the expression case to the value case. + *) +Lemma type_interp_mono : ∀ '(k, (A, ve)) δ k', k' ≤ k → type_interp ve A δ k → type_interp ve A δ k'. +Proof. + eapply (well_founded_ind (R := term_rel) term_rel_wf). + intros (k & A & []) IH δ k'; first last. + { (* expression rel *) + intros Hk He. simp type_interp in He. simp type_interp. intros e' n Hn Hred. + destruct (He e' n ltac:(lia) Hred) as (v & Hval & Hv). + exists v. split; first done. + eapply (IH (k-n, (A, inj_val v))); [ | lia | done]. + (* show that the induction is decreasing *) + dsimpl. + } + intros Hk Hv. + destruct A as [x | | | | A | A | A B | A B | A B | A ]; simp type_interp; simp type_interp in Hv. + - (* var case *) + by eapply sem_type_mono. + - (* universal case *) + destruct Hv as (e & -> & ? & Hv). + exists e. split_and!; [done.. | ]. intros τ. + eapply (IH (k, (A, inj_expr e))); [ dsimpl | done | done]. + - (* existential case *) + destruct Hv as (v' & -> & (τ & Hv)). exists v'. split; first done. + exists τ. eapply (IH (k, (A, _))); [ dsimpl | done..]. + - (* fun case *) + destruct Hv as (x & e & -> & ? & Hv). exists x, e. split_and!; [done..| ]. + intros v' k'' Hk'' Hv'. + apply Hv, Hv'. lia. + - (* pair case *) + destruct Hv as (v1 & v2 & -> & Hv1 & Hv2). + exists v1, v2. split_and!; first done. + all: eapply (IH (k, (_, _))); [ dsimpl | done..]. + - (* sum case *) + destruct Hv as [(v' & -> & Hv) | (v' & -> & Hv)]; [ left | right]. + all: exists v'; split; first done. + all: eapply (IH (k, (_, _))); [ dsimpl | done..]. + - (* rec case *) + destruct Hv as (v' & -> & ? & Hv). + exists v'. split_and!; [ done.. | ]. + intros k'' Hk''. apply Hv. lia. +Qed. + +(** We can now derive the two desired lemmas *) +Lemma val_rel_mono A δ k k' v : k' ≤ k → 𝒱 A δ k v → 𝒱 A δ k' v. +Proof. apply (type_interp_mono (k, (A, inj_val v))). Qed. +Lemma expr_rel_mono A δ k k' e : k' ≤ k → ℰ A δ k e → ℰ A δ k' e. +Proof. apply (type_interp_mono (k, (A, inj_expr e))). Qed. + + +(** Interpret a syntactic type *) +Program Definition interp_type A δ : sem_type := {| + sem_type_car := fun k v => 𝒱 A δ k v; +|}. +Next Obligation. by eapply val_rel_is_closed. Qed. +Next Obligation. by eapply val_rel_mono. Qed. + +(** Semantic typing of contexts *) +Implicit Types + (θ : gmap string expr). + +(** Context relation *) +Inductive sem_context_rel (δ : tyvar_interp) (k : nat) : typing_context → (gmap string expr) → Prop := + | sem_context_rel_empty : sem_context_rel δ k ∅ ∅ + | sem_context_rel_insert Γ θ v x A : + 𝒱 A δ k v → + sem_context_rel δ k Γ θ → + sem_context_rel δ k (<[x := A]> Γ) (<[x := of_val v]> θ). + +Notation 𝒢 := sem_context_rel. + + +(** *** Semantic typing judgment *) +Definition sem_typed Δ Γ e A := + is_closed (elements (dom Γ)) e ∧ + ∀ θ δ k, 𝒢 δ k Γ θ → ℰ A δ k (subst_map θ e). +Notation "'TY' Δ ; Γ ⊨ e : A" := (sem_typed Δ Γ e A) (at level 74, e, A at next level). + + +Lemma sem_context_rel_vals {δ k Γ θ x A} : + sem_context_rel δ k Γ θ → + Γ !! x = Some A → + ∃ e v, θ !! x = Some e ∧ to_val e = Some v ∧ 𝒱 A δ k v. +Proof. + induction 1 as [|Γ θ v y B Hvals Hctx IH]. + - naive_solver. + - rewrite lookup_insert_Some. intros [[-> ->]|[Hne Hlook]]. + + do 2 eexists. split; first by rewrite lookup_insert. + split; first by eapply to_of_val. done. + + eapply IH in Hlook as (e & w & Hlook & He & Hval). + do 2 eexists; split; first by rewrite lookup_insert_ne. + split; first done. done. +Qed. + +Lemma sem_context_rel_dom δ k Γ θ : + 𝒢 δ k Γ θ → dom Γ = dom θ. +Proof. + induction 1. + - by rewrite !dom_empty. + - rewrite !dom_insert. congruence. +Qed. + +Lemma sem_context_rel_dom_eq δ k Γ θ : + 𝒢 δ k Γ θ → dom Γ = dom θ. +Proof. + induction 1 as [ | ??????? IH]. + - rewrite !dom_empty //. + - rewrite !dom_insert IH //. +Qed. + +Lemma sem_context_rel_closed δ k Γ θ: + 𝒢 δ k Γ θ → subst_is_closed [] θ. +Proof. + induction 1 as [ | Γ θ v x A Hv Hctx IH]; rewrite /subst_is_closed. + - naive_solver. + - intros y e. rewrite lookup_insert_Some. + intros [[-> <-]|[Hne Hlook]]. + + by eapply val_rel_is_closed. + + eapply IH; last done. +Qed. + +Lemma sem_context_rel_mono Γ δ k k' θ : + k' ≤ k → 𝒢 δ k Γ θ → 𝒢 δ k' Γ θ. +Proof. + intros Hk. induction 1 as [|Γ θ v y B Hvals Hctx IH]; constructor. + - eapply val_rel_mono; done. + - apply IH. +Qed. + +Section boring_lemmas. + (** The lemmas in this section are all quite boring and expected statements, + but are quite technical to prove due to De Bruijn binders. + We encourage to skip over the proofs of these lemmas. + *) + + Lemma type_interp_ext : + ∀ '(k, (B, ve)), ∀ δ δ', + (∀ n k v, δ n k v ↔ δ' n k v) → + type_interp ve B δ k ↔ type_interp ve B δ' k. + Proof. + eapply (well_founded_ind (R := term_rel) term_rel_wf). + intros (k & A & [v|e]) IH δ δ'; first last. + { (* expression rel *) + intros Hd. simp type_interp. eapply forall_proper; intros e'. + eapply forall_proper; intros n. eapply if_iff; first done. + eapply if_iff; first done. f_equiv. intros v. f_equiv. + eapply (IH ((k - n), (A, inj_val v))); last done. + (* show that the induction is decreasing *) + dsimpl. + } + intros Hd. destruct A as [x | | | | A | A | A B | A B | A B | A ]; simp type_interp; eauto. + - f_equiv; intros e. f_equiv. f_equiv. + eapply forall_proper; intros τ. + eapply (IH (_, (_, _))); first dsimpl. + intros [|m] ?; simpl; eauto. + - f_equiv; intros w. f_equiv. f_equiv. intros τ. + eapply (IH (_, (_, _))); first dsimpl. + intros [|m] ?; simpl; eauto. + - f_equiv. intros ?. f_equiv. intros ?. + f_equiv. f_equiv. eapply forall_proper. intros ?. + eapply forall_proper. intros ?. + eapply forall_proper. intros ?. + eapply if_iff; by eapply (IH (_, (_, _))); first dsimpl. + - f_equiv. intros ?. f_equiv. intros ?. + f_equiv. f_equiv; by eapply (IH (_, (_, _))); first dsimpl. + - f_equiv; f_equiv; intros ?; f_equiv; by eapply (IH (_, (_, _))); first dsimpl. + - f_equiv; intros ?. f_equiv. f_equiv. + eapply forall_proper; intros ?. + eapply forall_proper; intros ?. + by eapply (IH (_, (_, _))); first dsimpl. + Qed. + + + Lemma type_interp_move_ren : + ∀ '(k, (B, ve)), ∀ δ σ, type_interp ve B (λ n, δ (σ n)) k ↔ type_interp ve (rename σ B) δ k. + Proof. + eapply (well_founded_ind (R := term_rel) term_rel_wf). + intros (k & A & [v|e]) IH δ σ; first last. + { (* expression rel *) + simp type_interp. eapply forall_proper; intros e'. + eapply forall_proper; intros n. eapply if_iff; first done. + eapply if_iff; first done. f_equiv. intros v. f_equiv. + eapply (IH (_, (_, _))). + (* show that the induction is decreasing *) + dsimpl. + } + destruct A as [x | | | | A | A | A B | A B | A B | A ]; simpl; simp type_interp; eauto. + - f_equiv; intros e. f_equiv. f_equiv. + eapply forall_proper; intros τ. + etransitivity; last eapply (IH (_, (_, _))); last dsimpl. + eapply (type_interp_ext (_, (_, _))). + intros [|m] ?; simpl; eauto. + - f_equiv; intros w. f_equiv. f_equiv. intros τ. + etransitivity; last eapply (IH (_, (_, _))); last dsimpl. + eapply (type_interp_ext (_, (_, _))). + intros [|m] ?; simpl; eauto. + - f_equiv. intros ?. f_equiv. intros ?. + f_equiv. f_equiv. eapply forall_proper. intros ?. + eapply forall_proper. intros ?. + eapply forall_proper. intros ?. + eapply if_iff; by eapply (IH (_, (_, _))); first dsimpl. + - f_equiv. intros ?. f_equiv. intros ?. + f_equiv. f_equiv; by eapply (IH (_, (_, _))); first dsimpl. + - f_equiv; f_equiv; intros ?; f_equiv; by eapply (IH (_, (_, _))); first dsimpl. + - f_equiv; intros ?. f_equiv. f_equiv. + eapply forall_proper; intros ?. + eapply forall_proper; intros ?. + etransitivity; first eapply (IH (_, (_, _))); first dsimpl. + (* NOTE: nice use of asimpl; :-) *) + asimpl. done. + Qed. + + + Lemma type_interp_move_subst : + ∀ '(k, (B, ve)), ∀ δ σ, type_interp ve B (λ n, interp_type (σ n) δ) k ↔ type_interp ve (B.[σ]) δ k. + Proof. + eapply (well_founded_ind (R := term_rel) term_rel_wf). + intros (k & A & [v|e]) IH δ σ; first last. + { (* expression rel *) + simp type_interp. eapply forall_proper; intros e'. + eapply forall_proper; intros n. eapply if_iff; first done. + eapply if_iff; first done. f_equiv. intros v. f_equiv. + eapply (IH (_, (_, _))). + (* show that the induction is decreasing *) + dsimpl. + } + destruct A as [x | | | | A | A | A B | A B | A B | A ]; simpl; simp type_interp; eauto. + - f_equiv; intros e. f_equiv. f_equiv. + eapply forall_proper; intros τ. + etransitivity; last eapply (IH (_, (_, _))); last dsimpl. + eapply (type_interp_ext (_, (_, _))). + intros [|m] ? ?; simpl. + + asimpl. simp type_interp. done. + + unfold up; simpl. etransitivity; + last eapply (type_interp_move_ren (_, (_, _))). + done. + - f_equiv; intros w. f_equiv. f_equiv. intros τ. + etransitivity; last eapply (IH (_, (_, _))); last dsimpl. + eapply (type_interp_ext (_, (_, _))). + intros [|m] ? ?; simpl. + + asimpl. simp type_interp. done. + + unfold up; simpl. etransitivity; + last eapply (type_interp_move_ren (_, (_, _))). + done. + - f_equiv. intros ?. f_equiv. intros ?. + f_equiv. f_equiv. eapply forall_proper. intros ?. + eapply forall_proper. intros ?. + eapply forall_proper. intros ?. + eapply if_iff; by eapply (IH (_, (_, _))); first dsimpl. + - f_equiv. intros ?. f_equiv. intros ?. + f_equiv. f_equiv; by eapply (IH (_, (_, _))); first dsimpl. + - f_equiv; f_equiv; intros ?; f_equiv; by eapply (IH (_, (_, _))); first dsimpl. + - f_equiv; intros ?. f_equiv. f_equiv. + eapply forall_proper; intros ?. + eapply forall_proper; intros ?. + etransitivity; first eapply (IH (_, (_, _))); first dsimpl. + (* NOTE: nice use of asimpl; :-) *) + asimpl. done. + Qed. + + + Lemma sem_val_rel_move_single_subst A B δ k v : + 𝒱 B (interp_type A δ .: δ) k v ↔ 𝒱 (B.[A/]) δ k v. + Proof. + etransitivity; last eapply (type_interp_move_subst (_, (_, _))). + eapply (type_interp_ext (_, (_, _))). + intros [| n] ? w; simpl; simp type_interp; done. + Qed. + + Lemma sem_expr_rel_move_single_subst A B δ k e : + ℰ B (interp_type A δ .: δ) k e ↔ ℰ (B.[A/]) δ k e. + Proof. + etransitivity; last eapply (type_interp_move_subst (_, (_, _))). + eapply (type_interp_ext (_, (_, _))). + intros [| n] ? w; simpl; simp type_interp; done. + Qed. + + Lemma sem_val_rel_cons A δ k v τ : + 𝒱 A δ k v ↔ 𝒱 A.[ren (+1)] (τ .: δ) k v. + Proof. + etransitivity; last eapply (type_interp_move_subst (_, (_, _))). + eapply (type_interp_ext (_, (_, _))). + intros [| n] ? w; simpl; simp type_interp; done. + Qed. + + Lemma sem_expr_rel_cons A δ k e τ : + ℰ A δ k e ↔ ℰ A.[ren (+1)] (τ .: δ) k e. + Proof. + etransitivity; last eapply (type_interp_move_subst (_, (_, _))). + eapply (type_interp_ext (_, (_, _))). + intros [| n] ? w; simpl; simp type_interp; done. + Qed. + + + Lemma sem_context_rel_cons Γ k δ θ τ : + 𝒢 δ k Γ θ → + 𝒢 (τ .: δ) k (⤉ Γ) θ. + Proof. + induction 1 as [ | Γ θ v x A Hv Hctx IH]; simpl. + - rewrite fmap_empty. constructor. + - rewrite fmap_insert. constructor; last done. + rewrite -sem_val_rel_cons. done. + Qed. +End boring_lemmas. + +(** Bind lemma *) +Lemma bind K e k δ A B : + ℰ A δ k e → + (∀ j v, j ≤ k → 𝒱 A δ j v → ℰ B δ j (fill K (of_val v))) → + ℰ B δ k (fill K e). +Proof. + intros H1 H2. simp type_interp in H1. simp type_interp. + intros e' n Hn (j & e'' & Hj & Hred1 & Hred2)%red_nsteps_fill. + specialize (H1 e'' j ltac:(lia) Hred1) as (v & Hev & Hv). + specialize (H2 (k-j) v ltac:(lia) Hv). + simp type_interp in H2. + rewrite (of_to_val _ _ Hev) in H2. + eapply H2 in Hred2; last lia. + assert (k - n = k - j - (n - j)) as -> by lia. + done. +Qed. + +(** This is the closure-under-expansion lemma from the lecture notes *) +Lemma expr_det_step_closure e e' A δ k : + det_step e e' → + ℰ A δ (k - 1) e' → + ℰ A δ k e. +Proof. + simp type_interp. intros Hdet Hexpr e'' n Hn [? Hred]%(det_step_red _ e'); last done. + destruct (Hexpr e'' (n -1)) as (v & Hev & Hv); [lia | done | ]. + exists v. split; first done. replace (k - n) with (k - 1 - (n - 1)) by lia. done. +Qed. + +Lemma expr_det_steps_closure e e' A δ k n : + nsteps det_step n e e' → ℰ A δ (k - n) e' → ℰ A δ k e. +Proof. + induction 1 as [ | n e1 e2 e3 Hstep Hsteps IH] in k |-* . + - replace (k - 0) with k by lia. done. + - intros He. + eapply expr_det_step_closure; first done. + apply IH. replace (k - 1 - n) with (k - (S n)) by lia. done. +Qed. + +(** ** Compatibility lemmas *) + +Lemma compat_int Δ Γ z : TY Δ; Γ ⊨ (Lit $ LitInt z) : Int. +Proof. + split; first done. + intros θ δ k _. + eapply (sem_val_expr_rel _ _ _ #z). + simp type_interp. eauto. +Qed. + +Lemma compat_bool Δ Γ b : TY Δ; Γ ⊨ (Lit $ LitBool b) : Bool. +Proof. + split; first done. + intros θ δ k _. + eapply (sem_val_expr_rel _ _ _ #b). simp type_interp. eauto. +Qed. + +Lemma compat_unit Δ Γ : TY Δ; Γ ⊨ (Lit $ LitUnit) : Unit. +Proof. + split; first done. + intros θ δ k _. + eapply (sem_val_expr_rel _ _ _ #LitUnit). simp type_interp. eauto. +Qed. + +Lemma compat_var Δ Γ x A : + Γ !! x = Some A → + TY Δ; Γ ⊨ (Var x) : A. +Proof. + intros Hx. split. + { eapply bool_decide_pack, elem_of_elements, elem_of_dom_2, Hx. } + intros θ δ k Hctx; simpl. + specialize (sem_context_rel_vals Hctx Hx) as (e & v & He & Heq & Hv). + rewrite He. simp type_interp. + rewrite -(of_to_val _ _ Heq). + intros e' n Hn [-> ->]%nsteps_val_inv. + rewrite to_of_val. eexists; split; first done. + replace (k -0) with k by lia. done. +Qed. + +Lemma compat_app Δ Γ e1 e2 A B : + TY Δ; Γ ⊨ e1 : (A → B) → + TY Δ; Γ ⊨ e2 : A → + TY Δ; Γ ⊨ (e1 e2) : B. +Proof. + intros [Hfuncl Hfun] [Hargcl Harg]. split. + { simpl. eauto. } + intros θ δ k Hctx; simpl. + specialize (Hfun _ _ _ Hctx). + specialize (Harg _ _ _ Hctx). + + eapply (bind [AppRCtx _]); first done. + intros j v Hj Hv. simpl. + + eapply (bind [AppLCtx _ ]). + { eapply expr_rel_mono; last done. lia. } + intros j' f Hj' Hf. + + simp type_interp in Hf. destruct Hf as (x & e & -> & Hcl & Hf). + specialize (Hf v j' (ltac:(lia))). + eapply expr_det_step_closure. + { eapply det_step_beta. apply is_val_of_val. } + eapply expr_rel_mono; last apply Hf; first lia. + eapply val_rel_mono; last done. lia. +Qed. + +Lemma is_closed_subst_map_delete X Γ (x: string) θ A e: + closed X e → + subst_is_closed [] θ → + dom Γ = dom θ → + (∀ y : string, y ∈ X → y ∈ dom (<[x:=A]> Γ)) → + is_closed (x :b: []) (subst_map (delete x θ) e). +Proof. + intros He Hθ Hdom1 Hdom2. + eapply closed_subst_weaken; [ | | apply He]. + - eapply subst_is_closed_subseteq; last done. + apply map_delete_subseteq. + - intros y Hy%Hdom2 Hn. apply elem_of_list_singleton. + apply not_elem_of_dom in Hn. apply elem_of_dom in Hy. + destruct (decide (x = y)) as [<- | Hneq]; first done. + rewrite lookup_delete_ne in Hn; last done. + rewrite lookup_insert_ne in Hy; last done. + move : Hn Hy. rewrite -elem_of_dom -not_elem_of_dom. + rewrite Hdom1. done. +Qed. + +(** Lambdas need to be closed by the context *) +Lemma compat_lam_named Δ Γ x e A B : + TY Δ; (<[ x := A ]> Γ) ⊨ e : B → + TY Δ; Γ ⊨ (Lam (BNamed x) e) : (A → B). +Proof. + intros [Hbodycl Hbody]. split. + { simpl. eapply is_closed_weaken; first eassumption. set_solver. } + intros θ δ k Hctxt. simpl. + eapply (sem_val_expr_rel _ _ _ (LamV x _)). + simp type_interp. + eexists (BNamed x), _. split_and!; [done| | ]. + { eapply is_closed_subst_map_delete; eauto. + + eapply sem_context_rel_closed in Hctxt. naive_solver. + + by eapply sem_context_rel_dom. + + apply elem_of_elements. + } + + intros v' k' Hk' Hv'. + specialize (Hbody (<[ x := of_val v']> θ) δ k'). + simpl. rewrite subst_subst_map. + 2: { by eapply sem_context_rel_closed. } + apply Hbody. + apply sem_context_rel_insert; first done. + eapply sem_context_rel_mono; last done. lia. +Qed. + +Lemma is_closed_subst_map_anon X Γ θ e: + closed X e → + subst_is_closed [] θ → + dom Γ = dom θ → + (∀ y, y ∈ X → y ∈ dom Γ) → + is_closed [] (subst_map θ e). +Proof. + intros He Hθ Hdom1 Hdom2. + eapply closed_subst_weaken; [ | | apply He]. + - eapply subst_is_closed_subseteq; done. + - intros y Hy%Hdom2 Hn. + apply not_elem_of_dom in Hn. apply elem_of_dom in Hy. + move : Hn Hy. rewrite -elem_of_dom -not_elem_of_dom. + rewrite Hdom1. done. +Qed. + +Lemma compat_lam_anon Δ Γ e A B : + TY Δ; Γ ⊨ e : B → + TY Δ; Γ ⊨ (Lam BAnon e) : (A → B). +Proof. + intros [Hbodycl Hbody]. split; first done. + intros θ δ k Hctxt. simpl. + eapply (sem_val_expr_rel _ _ _ (LamV BAnon _)). + simp type_interp. + eexists BAnon, _. split_and!; [done| | ]. + { eapply is_closed_subst_map_anon; eauto. + + eapply sem_context_rel_closed in Hctxt. naive_solver. + + by eapply sem_context_rel_dom. + + apply elem_of_elements. + } + + intros v' k' Hk' Hv'. + apply (Hbody θ δ k'). + eapply sem_context_rel_mono; last done. lia. +Qed. + +Lemma compat_int_binop Δ Γ op e1 e2 : + bin_op_typed op Int Int Int → + TY Δ; Γ ⊨ e1 : Int → + TY Δ; Γ ⊨ e2 : Int → + TY Δ; Γ ⊨ (BinOp op e1 e2) : Int. +Proof. + intros Hop [He1cl He1] [He2cl He2]. + split; first naive_solver. + + intros θ δ k Hctx. simpl. + specialize (He1 _ _ _ Hctx). + specialize (He2 _ _ _ Hctx). + + eapply (bind [BinOpRCtx _ _]); first done. + intros j v2 Hj Hv2. simpl. + + eapply (bind [BinOpLCtx _ _ ]). + { eapply expr_rel_mono; last done. lia. } + intros j' v1 Hj' Hv1. + + simp type_interp. intros e' n Hn Hred. + simp type_interp in Hv1. simp type_interp in Hv2. + destruct Hv1 as (z1 & ->). destruct Hv2 as (z2 & ->). + + inversion Hop; subst; simpl. + all: eapply det_step_red in Hred as [ ? Hred]; [ | eapply det_step_binop; done]. + all: apply nsteps_val_inv in Hred as [? ->]. + all: eexists; simpl; split; first done. + all: simp type_interp; eauto. +Qed. + +Lemma compat_int_bool_binop Δ Γ op e1 e2 : + bin_op_typed op Int Int Bool → + TY Δ; Γ ⊨ e1 : Int → + TY Δ; Γ ⊨ e2 : Int → + TY Δ; Γ ⊨ (BinOp op e1 e2) : Bool. +Proof. + intros Hop [He1cl He1] [He2cl He2]. + split; first naive_solver. + + intros θ δ k Hctx. simpl. + specialize (He1 _ _ _ Hctx). + specialize (He2 _ _ _ Hctx). + + eapply (bind [BinOpRCtx _ _]); first done. + intros j v2 Hj Hv2. simpl. + + eapply (bind [BinOpLCtx _ _ ]). + { eapply expr_rel_mono; last done. lia. } + intros j' v1 Hj' Hv1. + + simp type_interp. intros e' n Hn Hred. + simp type_interp in Hv1. simp type_interp in Hv2. + destruct Hv1 as (z1 & ->). destruct Hv2 as (z2 & ->). + + inversion Hop; subst; simpl. + all: eapply det_step_red in Hred as [ ? Hred]; [ | eapply det_step_binop; done]. + all: apply nsteps_val_inv in Hred as [? ->]. + all: eexists; simpl; split; first done. + all: simp type_interp; eauto. +Qed. + +Lemma compat_unop Δ Γ op A B e : + un_op_typed op A B → + TY Δ; Γ ⊨ e : A → + TY Δ; Γ ⊨ (UnOp op e) : B. +Proof. + intros Hop [Hecl He]. + split; first naive_solver. + + intros θ δ k Hctx. simpl. + specialize (He _ _ _ Hctx). + + eapply (bind [UnOpCtx _]); first done. + intros j v Hj Hv. simpl. + + simp type_interp. intros e' n Hn Hred. + inversion Hop; subst. + all: simp type_interp in Hv; destruct Hv as (? & ->). + all: eapply det_step_red in Hred as [ ? Hred]; [ | eapply det_step_unop; done]. + all: apply nsteps_val_inv in Hred as [? ->]. + all: eexists; simpl; split; first done. + all: simp type_interp; eauto. +Qed. + +Lemma compat_tlam Δ Γ e A : + TY S Δ; (⤉ Γ) ⊨ e : A → + TY Δ; Γ ⊨ (Λ, e) : (∀: A). +Proof. + intros [Hecl He]. split. + { simpl. by erewrite <-dom_fmap. } + + intros θ δ k Hctx. simpl. + simp type_interp. + intros e' n Hn Hred. eapply nsteps_val_inv' in Hred as [ -> ->]; last done. + eexists; split; first done. + simp type_interp. + eexists _. split_and!; [ done | | ]. + { eapply is_closed_subst_map_anon; eauto. + + eapply sem_context_rel_closed in Hctx; naive_solver. + + by eapply sem_context_rel_dom. + + rewrite dom_fmap. apply elem_of_elements. + } + + intros τ. eapply He. + replace (k -0) with k by lia. by eapply sem_context_rel_cons. +Qed. + +Lemma compat_tapp Δ Γ e A B : + type_wf Δ B → + TY Δ; Γ ⊨ e : (∀: A) → + TY Δ; Γ ⊨ (e <>) : (A.[B/]). +Proof. + intros Hwf [Hecl He]. + split; first naive_solver. + + intros θ δ k Hctx. simpl. + specialize (He _ _ _ Hctx). + eapply (bind [TAppCtx]); first done. + intros j v Hj Hv. + simp type_interp in Hv. + destruct Hv as (e' & -> & ? & He'). + + set (τ := interp_type B δ). + specialize (He' τ). simpl. + eapply expr_det_step_closure. + { apply det_step_tbeta. } + eapply sem_expr_rel_move_single_subst. + eapply expr_rel_mono; last done. + lia. +Qed. + +Lemma compat_pack Δ Γ e n A B : + type_wf n B → + type_wf (S n) A → + TY n; Γ ⊨ e : A.[B/] → + TY n; Γ ⊨ (pack e) : (∃: A). +Proof. + intros Hwf1 Hwf2 [Hecl He]. + split; first naive_solver. + + intros θ δ k Hctx. simpl. + + specialize (He _ _ _ Hctx). + eapply (bind [PackCtx]); first done. + intros j v Hj Hv. + simpl. eapply (sem_val_expr_rel _ _ _ (PackV v)). + simp type_interp. exists v; split; first done. + exists (interp_type B δ). + apply sem_val_rel_move_single_subst. done. +Qed. + +Lemma compat_unpack n Γ A B e e' x : + type_wf n B → + TY n; Γ ⊨ e : (∃: A) → + TY S n; <[x:=A]> (⤉Γ) ⊨ e' : B.[ren (+1)] → + TY n; Γ ⊨ (unpack e as BNamed x in e') : B. +Proof. + intros Hwf [Hecl He] [He'cl He']. split. + { simpl. apply andb_True. split; first done. eapply is_closed_weaken; first eassumption. set_solver. } + intros θ δ k Hctx. simpl. + + specialize (He _ _ _ Hctx). + eapply (bind [UnpackCtx _ _]); first done. + intros j v Hj Hv. + simp type_interp in Hv. destruct Hv as (v' & -> & τ & Hv'). + simpl. + + eapply expr_det_step_closure. + { apply det_step_unpack. apply is_val_of_val. } + simpl. rewrite subst_subst_map; last by eapply sem_context_rel_closed. + + specialize (He' (<[x := of_val v']> θ) (τ.:δ) (j - 1)). + rewrite <-sem_expr_rel_cons in He'. + apply He'. + constructor. + { eapply val_rel_mono; last done. lia. } + apply sem_context_rel_cons. + eapply sem_context_rel_mono; last done. lia. +Qed. + +Lemma compat_if n Γ e0 e1 e2 A : + TY n; Γ ⊨ e0 : Bool → + TY n; Γ ⊨ e1 : A → + TY n; Γ ⊨ e2 : A → + TY n; Γ ⊨ (if: e0 then e1 else e2) : A. +Proof. + intros [He0cl He0] [He1cl He1] [He2cl He2]. + split; first naive_solver. + + intros θ δ k Hctx. simpl. + specialize (He0 _ _ _ Hctx). + specialize (He1 _ _ _ Hctx). + specialize (He2 _ _ _ Hctx). + + eapply (bind [IfCtx _ _]); first done. + intros j v Hj Hv. + simp type_interp in Hv. destruct Hv as (b & ->). + simpl. + + destruct b. + - eapply expr_det_step_closure. + { apply det_step_if_true. } + eapply expr_rel_mono; last done. lia. + - eapply expr_det_step_closure. + { apply det_step_if_false. } + eapply expr_rel_mono; last done. lia. +Qed. + +Lemma compat_pair Δ Γ e1 e2 A B : + TY Δ; Γ ⊨ e1 : A → + TY Δ; Γ ⊨ e2 : B → + TY Δ; Γ ⊨ (e1, e2) : A × B. +Proof. + intros [He1cl He1] [He2cl He2]. + split; first naive_solver. + + intros θ δ k Hctx. simpl. + specialize (He1 _ _ _ Hctx). + specialize (He2 _ _ _ Hctx). + + eapply (bind [PairRCtx _]); first done. + intros j v2 Hj Hv2. + eapply (bind [PairLCtx _]). + { eapply expr_rel_mono; last done. lia. } + intros j' v1 Hj' Hv1. + + simpl. + eapply (sem_val_expr_rel _ _ _ (v1, v2)%V). + simp type_interp. exists v1, v2. split_and!; first done. + - done. + - eapply val_rel_mono; last done. lia. +Qed. + +Lemma compat_fst Δ Γ e A B : + TY Δ; Γ ⊨ e : A × B → + TY Δ; Γ ⊨ Fst e : A. +Proof. + intros [Hecl He]. split; first naive_solver. + + intros θ δ k Hctx. + specialize (He _ _ _ Hctx). + simpl. eapply (bind [FstCtx]); first done. + intros j v Hj Hv. + simp type_interp in Hv. destruct Hv as (v1 & v2 & -> & Hv1 & Hv2). + + eapply expr_det_step_closure. + { simpl. apply det_step_fst; apply is_val_of_val. } + eapply sem_val_expr_rel. eapply val_rel_mono; last done. lia. +Qed. + +Lemma compat_snd Δ Γ e A B : + TY Δ; Γ ⊨ e : A × B → + TY Δ; Γ ⊨ Snd e : B. +Proof. + intros [Hecl He]. split; first naive_solver. + + intros θ δ k Hctx. + specialize (He _ _ _ Hctx). + simpl. eapply (bind [SndCtx]); first done. + intros j v Hj Hv. + simp type_interp in Hv. destruct Hv as (v1 & v2 & -> & Hv1 & Hv2). + + eapply expr_det_step_closure. + { simpl. apply det_step_snd; apply is_val_of_val. } + eapply sem_val_expr_rel. eapply val_rel_mono; last done. lia. +Qed. + +Lemma compat_injl Δ Γ e A B : + TY Δ; Γ ⊨ e : A → + TY Δ; Γ ⊨ InjL e : A + B. +Proof. + (* TODO: exercise *) +Admitted. + + +Lemma compat_injr Δ Γ e A B : + TY Δ; Γ ⊨ e : B → + TY Δ; Γ ⊨ InjR e : A + B. +Proof. + (* TODO: exercise *) +Admitted. + + +Lemma compat_case Δ Γ e e1 e2 A B C : + TY Δ; Γ ⊨ e : B + C → + TY Δ; Γ ⊨ e1 : (B → A) → + TY Δ; Γ ⊨ e2 : (C → A) → + TY Δ; Γ ⊨ Case e e1 e2 : A. +Proof. + (* TODO: exercise *) +Admitted. + + +Lemma compat_roll n Γ e A : + TY n; Γ ⊨ e : (A.[(μ: A)%ty/]) → + TY n; Γ ⊨ (roll e) : (μ: A). +Proof. + intros [Hecl He]. split; first naive_solver. + + intros θ δ k Hctx. simpl. + specialize (He _ _ _ Hctx). + + eapply (bind [RollCtx]); first done. + intros j v Hj Hv. + eapply (sem_val_expr_rel _ _ _ (RollV v)). + + specialize (val_rel_is_closed _ _ _ _ Hv) as ?. + simp type_interp. + exists v. split_and!; [done.. | ]. + intros k' Hk'. eapply val_rel_mono; last done. lia. +Qed. + +Lemma compat_unroll n Γ e A : + TY n; Γ ⊨ e : (μ: A) → + TY n; Γ ⊨ (unroll e) : (A.[(μ: A)%ty/]). +Proof. + intros [Hecl He]. split; first naive_solver. + + intros θ δ k Hctx. simpl. + specialize (He _ _ _ Hctx). + + eapply (bind [UnrollCtx]); first done. + intros j v Hj Hv. + destruct j as [ | j]; first by apply sem_expr_rel_zero_trivial. + simp type_interp in Hv. destruct Hv as (v' & -> & ? & Hv). + eapply expr_det_step_closure. + { simpl. apply det_step_unroll. apply is_val_of_val. } + eapply sem_val_expr_rel. apply Hv. lia. +Qed. + +Local Hint Resolve compat_var compat_lam_named compat_lam_anon compat_tlam compat_int compat_bool compat_unit compat_if compat_app compat_tapp compat_pack compat_unpack compat_int_binop compat_int_bool_binop compat_unop compat_pair compat_fst compat_snd compat_injl compat_injr compat_case compat_roll compat_unroll: core. + +Lemma sem_soundness Δ Γ e A : + TY Δ; Γ ⊢ e : A → + TY Δ; Γ ⊨ e : A. +Proof. + induction 1 as [ | Δ Γ x e A B Hsyn IH | Δ Γ e A B Hsyn IH| Δ Γ e A Hsyn IH| | | | | | | | | n Γ e1 e2 op A B C [] ? ? ? ? | | | | | | | | | ]; eauto. +Qed. + + +(* dummy delta which we can use if we don't care *) +Program Definition any_type : sem_type := {| sem_type_car := λ k v, is_closed [] v |}. +Definition δ_any : var → sem_type := λ _, any_type. + + +Definition safe e := + ∀ e' n, red_nsteps n e e' → is_val e'. + +Lemma type_safety e A : + TY 0; ∅ ⊢ e : A → + safe e. +Proof. + intros [Hecl He]%sem_soundness e' n Hred. + specialize (He ∅ δ_any (S n)). simp type_interp in He. + rewrite subst_map_empty in He. + edestruct (He ) as (v & Hv & _); [ | | eassumption | ]. + - constructor. + - lia. + - rewrite <- (of_to_val _ _ Hv). apply is_val_of_val. +Qed. + + +(** Additional lemmas *) +Lemma semantic_app A B δ k e1 e2 : + ℰ (A → B) δ k e1 → + ℰ A δ k e2 → + ℰ B δ k (e1 e2). +Proof. + intros He1 He2. + eapply (bind [AppRCtx e1]); first done. + intros j v Hj Hv. eapply (bind [AppLCtx _]). + { eapply expr_rel_mono; last done. lia. } + intros j' v' Hj' Hf. + simp type_interp in Hf. destruct Hf as (x & e & -> & Hcl & Hf). + eapply expr_det_step_closure. + { apply det_step_beta. apply is_val_of_val. } + apply Hf; first lia. + eapply val_rel_mono; last done. lia. +Qed. diff --git a/theories/type_systems/systemf_mu/notation.v b/theories/type_systems/systemf_mu/notation.v new file mode 100644 index 0000000..d0398d8 --- /dev/null +++ b/theories/type_systems/systemf_mu/notation.v @@ -0,0 +1,97 @@ +From semantics.ts.systemf_mu Require Export lang. +Set Default Proof Using "Type". + + +(** Coercions to make programs easier to type. *) +Coercion of_val : val >-> expr. +Coercion LitInt : Z >-> base_lit. +Coercion LitBool : bool >-> base_lit. +Coercion App : expr >-> Funclass. +Coercion Var : string >-> expr. + +(** Define some derived forms. *) +Notation Let x e1 e2 := (App (Lam x e2) e1) (only parsing). +Notation Seq e1 e2 := (Let BAnon e1 e2) (only parsing). +Notation Match e0 x1 e1 x2 e2 := (Case e0 (Lam x1 e1) (Lam x2 e2)) (only parsing). + +(* No scope for the values, does not conflict and scope is often not inferred +properly. *) +Notation "# l" := (LitV l%Z%V%stdpp) (at level 8, format "# l"). +Notation "# l" := (Lit l%Z%E%stdpp) (at level 8, format "# l") : expr_scope. + +(** Syntax inspired by Coq/Ocaml. Constructions with higher precedence come + first. *) +Notation "( e1 , e2 , .. , en )" := (Pair .. (Pair e1 e2) .. en) : expr_scope. +Notation "( e1 , e2 , .. , en )" := (PairV .. (PairV e1 e2) .. en) : val_scope. + +Notation "'match:' e0 'with' 'InjL' x1 => e1 | 'InjR' x2 => e2 'end'" := + (Match e0 x1%binder e1 x2%binder e2) + (e0, x1, e1, x2, e2 at level 200, + format "'[hv' 'match:' e0 'with' '/ ' '[' 'InjL' x1 => '/ ' e1 ']' '/' '[' | 'InjR' x2 => '/ ' e2 ']' '/' 'end' ']'") : expr_scope. +Notation "'match:' e0 'with' 'InjR' x1 => e1 | 'InjL' x2 => e2 'end'" := + (Match e0 x2%binder e2 x1%binder e1) + (e0, x1, e1, x2, e2 at level 200, only parsing) : expr_scope. + +Notation "()" := LitUnit : val_scope. + +Notation "e1 + e2" := (BinOp PlusOp e1%E e2%E) : expr_scope. +Notation "e1 - e2" := (BinOp MinusOp e1%E e2%E) : expr_scope. +Notation "e1 * e2" := (BinOp MultOp e1%E e2%E) : expr_scope. +Notation "e1 ≤ e2" := (BinOp LeOp e1%E e2%E) : expr_scope. +Notation "e1 = e2" := (BinOp EqOp e1%E e2%E) : expr_scope. +Notation "e1 < e2" := (BinOp LtOp e1%E e2%E) : expr_scope. + +(*Notation "~ e" := (UnOp NegOp e%E) (at level 75, right associativity) : expr_scope.*) + +Notation "'if:' e1 'then' e2 'else' e3" := (If e1%E e2%E e3%E) + (at level 200, e1, e2, e3 at level 200) : expr_scope. + +Notation "λ: x , e" := (Lam x%binder e%E) + (at level 200, x at level 1, e at level 200, + format "'[' 'λ:' x , '/ ' e ']'") : expr_scope. +Notation "λ: x y .. z , e" := (Lam x%binder (Lam y%binder .. (Lam z%binder e%E) ..)) + (at level 200, x, y, z at level 1, e at level 200, + format "'[' 'λ:' x y .. z , '/ ' e ']'") : expr_scope. + +Notation "λ: x , e" := (LamV x%binder e%E) + (at level 200, x at level 1, e at level 200, + format "'[' 'λ:' x , '/ ' e ']'") : val_scope. +Notation "λ: x y .. z , e" := (LamV x%binder (Lam y%binder .. (Lam z%binder e%E) .. )) + (at level 200, x, y, z at level 1, e at level 200, + format "'[' 'λ:' x y .. z , '/ ' e ']'") : val_scope. + +Notation "'let:' x := e1 'in' e2" := (Lam x%binder e2%E e1%E) + (at level 200, x at level 1, e1, e2 at level 200, + format "'[' 'let:' x := '[' e1 ']' 'in' '/' e2 ']'") : expr_scope. +Notation "e1 ;; e2" := (Lam BAnon e2%E e1%E) + (at level 100, e2 at level 200, + format "'[' '[hv' '[' e1 ']' ;; ']' '/' e2 ']'") : expr_scope. + + +Notation "'Λ' , e" := (TLam e%E) + (at level 200, e at level 200, + format "'[' 'Λ' , '/ ' e ']'") : expr_scope. +Notation "'Λ' , e" := (TLamV e%E) + (at level 200, e at level 200, + format "'[' 'Λ' , '/ ' e ']'") : val_scope. + +(* the [e] always needs to be paranthesized, due to the level + (chosen to make this cooperate with the [App] coercion) *) +Notation "e '<>'" := (TApp e%E) + (at level 10) : expr_scope. +(*Check ((Λ, #4) <>)%E.*) +(*Check (((λ: "x", "x") #5) <>)%E.*) + +Notation "'pack' e" := (Pack e%E) + (at level 200, e at level 200) : expr_scope. +Notation "'pack' v" := (PackV v%V) + (at level 200, v at level 200) : val_scope. +Notation "'unpack' e1 'as' x 'in' e2" := (Unpack x%binder e1%E e2%E) + (at level 200, e1, e2 at level 200, x at level 1) : expr_scope. + +Notation "'roll' e" := (Roll e%E) + (at level 200, e at level 200) : expr_scope. +Notation "'roll' v" := (RollV v%E) + (at level 200, v at level 200) : val_scope. +Notation "'unroll' e" := (Unroll e%E) + (at level 200, e at level 200) : expr_scope. diff --git a/theories/type_systems/systemf_mu/parallel_subst.v b/theories/type_systems/systemf_mu/parallel_subst.v new file mode 100644 index 0000000..4961307 --- /dev/null +++ b/theories/type_systems/systemf_mu/parallel_subst.v @@ -0,0 +1,275 @@ +From stdpp Require Import prelude. +From iris Require Import prelude. +From semantics.lib Require Import facts maps. +From semantics.ts.systemf_mu Require Import lang. + +Fixpoint subst_map (xs : gmap string expr) (e : expr) : expr := + match e with + | Lit l => Lit l + | Var y => match xs !! y with Some es => es | _ => Var y end + | App e1 e2 => App (subst_map xs e1) (subst_map xs e2) + | Lam x e => Lam x (subst_map (binder_delete x xs) e) + | UnOp op e => UnOp op (subst_map xs e) + | BinOp op e1 e2 => BinOp op (subst_map xs e1) (subst_map xs e2) + | If e0 e1 e2 => If (subst_map xs e0) (subst_map xs e1) (subst_map xs e2) + | TApp e => TApp (subst_map xs e) + | TLam e => TLam (subst_map xs e) + | Pack e => Pack (subst_map xs e) + | Unpack x e1 e2 => Unpack x (subst_map xs e1) (subst_map (binder_delete x xs) e2) + | Pair e1 e2 => Pair (subst_map xs e1) (subst_map xs e2) + | Fst e => Fst (subst_map xs e) + | Snd e => Snd (subst_map xs e) + | InjL e => InjL (subst_map xs e) + | InjR e => InjR (subst_map xs e) + | Case e e1 e2 => Case (subst_map xs e) (subst_map xs e1) (subst_map xs e2) + | Roll e => Roll (subst_map xs e) + | Unroll e => Unroll (subst_map xs e) + end. + +Lemma subst_map_empty e : + subst_map ∅ e = e. +Proof. + induction e; simpl; f_equal; eauto. + all: destruct x; simpl; [done | by rewrite !delete_empty..]. +Qed. + +Lemma subst_map_is_closed X e xs : + is_closed X e → + (∀ x : string, x ∈ dom xs → x ∉ X) → + subst_map xs e = e. +Proof. + intros Hclosed Hd. + induction e in xs, X, Hd, Hclosed |-*; simpl in *;try done. + all: repeat match goal with + | H : Is_true (_ && _) |- _ => apply andb_True in H as [ ? ? ] + end. + { (* Var *) + apply bool_decide_spec in Hclosed. + assert (xs !! x = None) as ->; last done. + destruct (xs !! x) as [s | ] eqn:Helem; last done. + exfalso; eapply Hd; last apply Hclosed. + apply elem_of_dom; eauto. + } + { (* lambdas *) + erewrite IHe; [done | done |]. + intros y. destruct x as [ | x]; first apply Hd. + simpl. + rewrite dom_delete elem_of_difference not_elem_of_singleton. + intros [Hnx%Hd Hneq]. rewrite elem_of_cons. intros [? | ?]; done. + } + 8: { (* Unpack *) + erewrite IHe1; [ | done | done]. + erewrite IHe2; [ done | done | ]. + intros y. destruct x as [ | x]; first apply Hd. + simpl. + rewrite dom_delete elem_of_difference not_elem_of_singleton. + intros [Hnx%Hd Hneq]. rewrite elem_of_cons. intros [? | ?]; done. + } + (* all other cases *) + all: repeat match goal with + | H : ∀ _ _, _ → _ → subst_map _ _ = _ |- _ => erewrite H; clear H + end; done. +Qed. + +Lemma subst_map_subst map x (e e' : expr) : + is_closed [] e' → + subst_map map (subst x e' e) = subst_map (<[x:=e']>map) e. +Proof. + intros He'. + revert x map; induction e; intros xx map; simpl; + try (f_equal; eauto). + - case_decide. + + simplify_eq/=. rewrite lookup_insert. + rewrite (subst_map_is_closed []); [done | apply He' | ]. + intros ? ?. apply not_elem_of_nil. + + rewrite lookup_insert_ne; done. + - destruct x; simpl; first done. + + case_decide. + * simplify_eq/=. by rewrite delete_insert_delete. + * rewrite delete_insert_ne; last by congruence. done. + - destruct x; simpl; first done. + + case_decide. + * simplify_eq/=. by rewrite delete_insert_delete. + * rewrite delete_insert_ne; last by congruence. done. +Qed. + +Definition subst_is_closed (X : list string) (map : gmap string expr) := + ∀ x e, map !! x = Some e → closed X e. +Lemma subst_is_closed_subseteq X map1 map2 : + map1 ⊆ map2 → subst_is_closed X map2 → subst_is_closed X map1. +Proof. + intros Hsub Hclosed2 x e Hl. eapply Hclosed2, map_subseteq_spec; done. +Qed. + +Lemma subst_subst_map x es map e : + subst_is_closed [] map → + subst x es (subst_map (delete x map) e) = + subst_map (<[x:=es]>map) e. +Proof. + revert map es x; induction e; intros map v0 xx Hclosed; simpl; + try (f_equal; eauto). + - destruct (decide (xx=x)) as [->|Hne]. + + rewrite lookup_delete // lookup_insert //. simpl. + rewrite decide_True //. + + rewrite lookup_delete_ne // lookup_insert_ne //. + destruct (_ !! x) as [rr|] eqn:Helem. + * apply Hclosed in Helem. + by apply subst_is_closed_nil. + * simpl. rewrite decide_False //. + - destruct x; simpl; first by auto. + case_decide. + + simplify_eq. rewrite delete_idemp delete_insert_delete. done. + + rewrite delete_insert_ne //; last congruence. rewrite delete_commute. apply IHe. + eapply subst_is_closed_subseteq; last done. + apply map_delete_subseteq. + - destruct x; simpl; first by auto. + case_decide. + + simplify_eq. rewrite delete_idemp delete_insert_delete. done. + + rewrite delete_insert_ne //; last congruence. rewrite delete_commute. apply IHe2. + eapply subst_is_closed_subseteq; last done. + apply map_delete_subseteq. +Qed. + +Lemma subst'_subst_map b (es : expr) map e : + subst_is_closed [] map → + subst' b es (subst_map (binder_delete b map) e) = + subst_map (binder_insert b es map) e. +Proof. + destruct b; first done. + apply subst_subst_map. +Qed. + +Lemma closed_subst_weaken e map X Y : + subst_is_closed [] map → + (∀ x, x ∈ X → x ∉ dom map → x ∈ Y) → + closed X e → + closed Y (subst_map map e). +Proof. + induction e in X, Y, map |-*; rewrite /closed; simpl; intros Hmclosed Hsub Hcl; first done. + all: repeat match goal with + | H : Is_true (_ && _) |- _ => apply andb_True in H as [ ? ? ] + end. + { (* vars *) + destruct (map !! x) as [es | ] eqn:Heq. + + apply is_closed_weaken_nil. by eapply Hmclosed. + + apply bool_decide_pack. apply Hsub; first by eapply bool_decide_unpack. + by apply not_elem_of_dom. + } + { (* lambdas *) + eapply IHe; last done. + + eapply subst_is_closed_subseteq; last done. + destruct x; first done. apply map_delete_subseteq. + + intros y. destruct x as [ | x]; first by apply Hsub. + rewrite !elem_of_cons. intros [-> | Hy] Hn; first by left. + destruct (decide (y = x)) as [ -> | Hneq]; first by left. + right. eapply Hsub; first done. set_solver. + } + 8: { (* unpack *) + apply andb_True; split; first naive_solver. + eapply IHe2; last done. + + eapply subst_is_closed_subseteq; last done. + destruct x; first done. apply map_delete_subseteq. + + intros y. destruct x as [ | x]; first by apply Hsub. + rewrite !elem_of_cons. intros [-> | Hy] Hn; first by left. + destruct (decide (y = x)) as [ -> | Hneq]; first by left. + right. eapply Hsub; first done. set_solver. + } + (* all other cases *) + all: repeat match goal with + | |- Is_true (_ && _) => apply andb_True; split + end. + all: naive_solver. +Qed. + +Lemma subst_is_closed_weaken X1 X2 θ : + subst_is_closed X1 θ → + X1 ⊆ X2 → + subst_is_closed X2 θ. +Proof. + intros Hcl Hincl x e Hlook. + eapply is_closed_weaken; last done. + by eapply Hcl. +Qed. + +Lemma subst_is_closed_weaken_nil X θ : + subst_is_closed [] θ → + subst_is_closed X θ. +Proof. + intros; eapply subst_is_closed_weaken; first done. + apply list_subseteq_nil. +Qed. + +Lemma subst_is_closed_insert X e f θ : + is_closed X e → + subst_is_closed X (delete f θ) → + subst_is_closed X (<[f := e]> θ). +Proof. + intros Hcl Hcl' x e'. + destruct (decide (x = f)) as [<- | ?]. + - rewrite lookup_insert. intros [= <-]. done. + - rewrite lookup_insert_ne; last done. + intros Hlook. eapply Hcl'. + rewrite lookup_delete_ne; done. +Qed. + +(* NOTE: this is a simplified version of the tactic in tactics.v which + suffice for this proof *) +Ltac simplify_closed := + repeat match goal with + | |- closed _ _ => unfold closed; simpl + | |- Is_true (_ && _) => simpl; rewrite !andb_True; split_and! + | H : closed _ _ |- _ => unfold closed in H; simpl in H + | H: Is_true (_ && _) |- _ => simpl in H; apply andb_True in H + | H : _ ∧ _ |- _ => destruct H + end. + +Lemma is_closed_subst_map X θ e : + subst_is_closed X θ → + closed (X ++ elements (dom θ)) e → + closed X (subst_map θ e). +Proof. + induction e in X, θ |-*; eauto. + all: try solve [intros; simplify_closed; naive_solver]. + - unfold subst_map. + destruct (θ !! x) eqn:Heq. + + intros Hcl _. eapply Hcl; done. + + intros _ Hcl. + apply closed_is_closed in Hcl. + apply bool_decide_eq_true in Hcl. + apply elem_of_app in Hcl. + destruct Hcl as [ | H]. + * apply closed_is_closed. by apply bool_decide_eq_true. + * apply elem_of_elements in H. + by apply not_elem_of_dom in Heq. + - intros. simplify_closed. + eapply IHe. + + destruct x as [ | x]; simpl; first done. + intros y e'. rewrite lookup_delete_Some. intros (? & Hlook%H). + eapply is_closed_weaken; first done. + by apply list_subseteq_cons_r. + + eapply is_closed_weaken; first done. + simpl. destruct x as [ | x]; first done; simpl. + apply list_subseteq_cons_l. + apply stdpp.sets.elem_of_subseteq. + intros y; simpl. rewrite elem_of_cons !elem_of_app. + intros [ | Helem]; first naive_solver. + destruct (decide (x = y)) as [<- | Hneq]; first naive_solver. + right. right. apply elem_of_elements. rewrite dom_delete elem_of_difference elem_of_singleton. + apply elem_of_elements in Helem; done. + - intros. simplify_closed. { naive_solver. } + (* same proof as for lambda *) + eapply IHe2. + + destruct x as [ | x]; simpl; first done. + intros y e'. rewrite lookup_delete_Some. intros (? & Hlook%H). + eapply is_closed_weaken; first done. + by apply list_subseteq_cons_r. + + eapply is_closed_weaken; first done. + simpl. destruct x as [ | x]; first done; simpl. + apply list_subseteq_cons_l. + apply stdpp.sets.elem_of_subseteq. + intros y; simpl. rewrite elem_of_cons !elem_of_app. + intros [ | Helem]; first naive_solver. + destruct (decide (x = y)) as [<- | Hneq]; first naive_solver. + right. right. apply elem_of_elements. rewrite dom_delete elem_of_difference elem_of_singleton. + apply elem_of_elements in Helem; done. +Qed. diff --git a/theories/type_systems/systemf_mu/pure.v b/theories/type_systems/systemf_mu/pure.v new file mode 100644 index 0000000..d7c942a --- /dev/null +++ b/theories/type_systems/systemf_mu/pure.v @@ -0,0 +1,289 @@ +From stdpp Require Import gmap base relations. +From iris Require Import prelude. +From semantics.lib Require Import maps. +From semantics.ts.systemf_mu Require Import lang notation. + +Lemma contextual_ectx_step_case K e e' : + contextual_step (fill K e) e' → + (∃ e'', e' = fill K e'' ∧ contextual_step e e'') ∨ is_val e. +Proof. + (* TODO: exercise *) +Admitted. + + +(** ** Deterministic reduction *) + +Record det_step (e1 e2 : expr) := { + det_step_safe : reducible e1; + det_step_det e2' : + contextual_step e1 e2' → e2' = e2 +}. + +Record det_base_step (e1 e2 : expr) := { + det_base_step_safe : base_reducible e1; + det_base_step_det e2' : + base_step e1 e2' → e2' = e2 +}. + +Lemma det_base_step_det_step e1 e2 : det_base_step e1 e2 → det_step e1 e2. +Proof. + intros [Hp1 Hp2]. split. + - destruct Hp1 as (e2' & ?). + eexists e2'. by apply base_contextual_step. + - intros e2' ?%base_reducible_contextual_step; [ | done]. by apply Hp2. +Qed. + +(** *** Pure execution lemmas *) +Local Ltac inv_step := + repeat match goal with + | H : to_val _ = Some _ |- _ => apply of_to_val in H + | H : base_step ?e ?e2 |- _ => + try (is_var e; fail 1); (* inversion yields many goals if [e] is a variable + and should thus better be avoided. *) + inversion H; subst; clear H + end. +Local Ltac solve_exec_safe := intros; subst; eexists; econstructor; eauto. +Local Ltac solve_exec_detdet := simpl; intros; inv_step; try done. +Local Ltac solve_det_exec := + subst; intros; apply det_base_step_det_step; + constructor; [solve_exec_safe | solve_exec_detdet]. + +Lemma det_step_beta x e e2 : + is_val e2 → + det_step (App (@Lam x e) e2) (subst' x e2 e). +Proof. solve_det_exec. Qed. + +Lemma det_step_tbeta e : + det_step ((Λ, e) <>) e. +Proof. solve_det_exec. Qed. + +Lemma det_step_unpack e1 e2 x : + is_val e1 → + det_step (unpack (pack e1) as x in e2) (subst' x e1 e2). +Proof. solve_det_exec. Qed. + +Lemma det_step_unop op e v v' : + to_val e = Some v → + un_op_eval op v = Some v' → + det_step (UnOp op e) v'. +Proof. solve_det_exec. by simplify_eq. Qed. + +Lemma det_step_binop op e1 v1 e2 v2 v' : + to_val e1 = Some v1 → + to_val e2 = Some v2 → + bin_op_eval op v1 v2 = Some v' → + det_step (BinOp op e1 e2) v'. +Proof. solve_det_exec. by simplify_eq. Qed. + +Lemma det_step_if_true e1 e2 : + det_step (if: #true then e1 else e2) e1. +Proof. solve_det_exec. Qed. +Lemma det_step_if_false e1 e2 : + det_step (if: #false then e1 else e2) e2. +Proof. solve_det_exec. Qed. + +Lemma det_step_fst e1 e2 : + is_val e1 → + is_val e2 → + det_step (Fst (e1, e2)) e1. +Proof. solve_det_exec. Qed. +Lemma det_step_snd e1 e2 : + is_val e1 → + is_val e2 → + det_step (Snd (e1, e2)) e2. +Proof. solve_det_exec. Qed. + +Lemma det_step_casel e e1 e2 : + is_val e → + det_step (Case (InjL e) e1 e2) (e1 e). +Proof. solve_det_exec. Qed. +Lemma det_step_caser e e1 e2 : + is_val e → + det_step (Case (InjR e) e1 e2) (e2 e). +Proof. solve_det_exec. Qed. + +Lemma det_step_unroll e : + is_val e → + det_step (unroll (roll e)) e. +Proof. solve_det_exec. Qed. + +(** ** n-step reduction *) +(** Reduce in n steps to an irreducible expression. + (this is ⇝^n from the lecture notes) + *) +Definition red_nsteps (n : nat) (e e' : expr) := nsteps contextual_step n e e' ∧ irreducible e'. + +Lemma det_step_red e e' e'' n : + det_step e e' → + red_nsteps n e e'' → + 1 ≤ n ∧ red_nsteps (n - 1) e' e''. +Proof. + intros [Hprog Hstep] Hred. + inversion Hprog; subst. + destruct Hred as [Hred Hirred]. + destruct n as [ | n]. + { inversion Hred; subst. + exfalso; eapply not_reducible; done. + } + inversion Hred; subst. simpl. + apply Hstep in H as ->. apply Hstep in H1 as ->. + split; first lia. + replace (n - 0) with n by lia. done. +Qed. + +Lemma contextual_step_red_nsteps n e e' e'' : + contextual_step e e' → + red_nsteps n e' e'' → + red_nsteps (S n) e e''. +Proof. + intros Hstep [Hsteps Hirred]. + split; last done. + by econstructor. +Qed. + +Lemma nsteps_val_inv n v e' : + red_nsteps n (of_val v) e' → n = 0 ∧ e' = of_val v. +Proof. + intros [Hred Hirred]; cbn in *. + destruct n as [ | n]. + - inversion Hred; subst. done. + - inversion Hred; subst. exfalso. eapply val_irreducible; last done. + rewrite to_of_val. eauto. +Qed. + +Lemma nsteps_val_inv' n v e e' : + to_val e = Some v → + red_nsteps n e e' → n = 0 ∧ e' = of_val v. +Proof. intros Ht. rewrite -(of_to_val _ _ Ht). apply nsteps_val_inv. Qed. + +Lemma red_nsteps_fill K k e e' : + red_nsteps k (fill K e) e' → + ∃ j e'', j ≤ k ∧ + red_nsteps j e e'' ∧ + red_nsteps (k - j) (fill K e'') e'. +Proof. + (* TODO: exercise *) +Admitted. + + + +(** Additionally useful stepping lemmas *) +Lemma app_step_r (e1 e2 e2': expr) : + contextual_step e2 e2' → contextual_step (e1 e2) (e1 e2'). +Proof. by apply (fill_contextual_step [AppRCtx _]). Qed. + +Lemma app_step_l (e1 e1' e2: expr) : + contextual_step e1 e1' → is_val e2 → contextual_step (e1 e2) (e1' e2). +Proof. + intros ? (v & Hv)%is_val_spec. + rewrite <-(of_to_val _ _ Hv). + by apply (fill_contextual_step [AppLCtx _]). +Qed. + +Lemma app_step_beta (x: string) (e e': expr) : + is_val e' → is_closed [x] e → contextual_step ((λ: x, e) e') (lang.subst x e' e). +Proof. + intros Hval Hclosed. eapply base_contextual_step, BetaS; eauto. +Qed. + +Lemma unroll_roll_step (e: expr) : + is_val e → contextual_step (unroll (roll e)) e. +Proof. + intros ?; by eapply base_contextual_step, UnrollS. +Qed. + + +Lemma fill_reducible K e : + reducible e → reducible (fill K e). +Proof. + intros (e' & Hstep). + exists (fill K e'). eapply fill_contextual_step. done. +Qed. + +Lemma reducible_contextual_step_case K e e' : + contextual_step (fill K e) (e') → + reducible e → + ∃ e'', e' = fill K e'' ∧ contextual_step e e''. +Proof. + intros [ | Hval]%contextual_ectx_step_case Hred; first done. + exfalso. apply is_val_spec in Hval as (v & Hval). + apply reducible_not_val in Hred. congruence. +Qed. + +(** Contextual lifting lemmas for deterministic reduction *) +Tactic Notation "lift_det" uconstr(ctx) := + intros; + let Hs := fresh in + match goal with + | H : det_step _ _ |- _ => destruct H as [? Hs] + end; + simplify_val; econstructor; + [intros; by eapply (fill_reducible ctx) | + intros ? (? & -> & ->%Hs)%(reducible_contextual_step_case ctx); done ]. + +Lemma det_step_pair_r e1 e2 e2' : + det_step e2 e2' → + det_step (e1, e2)%E (e1, e2')%E. +Proof. lift_det [PairRCtx _]. Qed. +Lemma det_step_pair_l e1 e1' e2 : + is_val e2 → + det_step e1 e1' → + det_step (e1, e2)%E (e1', e2)%E. +Proof. lift_det [PairLCtx _]. Qed. +Lemma det_step_binop_r e1 e2 e2' op : + det_step e2 e2' → + det_step (BinOp op e1 e2)%E (BinOp op e1 e2')%E. +Proof. lift_det [BinOpRCtx _ _]. Qed. +Lemma det_step_binop_l e1 e1' e2 op : + is_val e2 → + det_step e1 e1' → + det_step (BinOp op e1 e2)%E (BinOp op e1' e2)%E. +Proof. lift_det [BinOpLCtx _ _]. Qed. +Lemma det_step_if e e' e1 e2 : + det_step e e' → + det_step (If e e1 e2)%E (If e' e1 e2)%E. +Proof. lift_det [IfCtx _ _]. Qed. +Lemma det_step_app_r e1 e2 e2' : + det_step e2 e2' → + det_step (App e1 e2)%E (App e1 e2')%E. +Proof. lift_det [AppRCtx _]. Qed. +Lemma det_step_app_l e1 e1' e2 : + is_val e2 → + det_step e1 e1' → + det_step (App e1 e2)%E (App e1' e2)%E. +Proof. lift_det [AppLCtx _]. Qed. +Lemma det_step_snd_lift e e' : + det_step e e' → + det_step (Snd e)%E (Snd e')%E. +Proof. lift_det [SndCtx]. Qed. +Lemma det_step_fst_lift e e' : + det_step e e' → + det_step (Fst e)%E (Fst e')%E. +Proof. lift_det [FstCtx]. Qed. + + +#[global] +Hint Resolve app_step_r app_step_l app_step_beta unroll_roll_step : core. +#[global] +Hint Extern 1 (is_val _) => (simpl; fast_done) : core. +#[global] +Hint Immediate is_val_of_val : core. + +#[global] +Hint Resolve det_step_beta det_step_tbeta det_step_unpack det_step_unop det_step_binop det_step_if_true det_step_if_false det_step_fst det_step_snd det_step_casel det_step_caser det_step_unroll : core. + +#[global] +Hint Resolve det_step_pair_r det_step_pair_l det_step_binop_r det_step_binop_l det_step_if det_step_app_r det_step_app_l det_step_snd_lift det_step_fst_lift : core. + +#[global] +Hint Constructors nsteps : core. + +#[global] +Hint Extern 1 (is_val _) => simpl : core. + +(** Prove a single deterministic step using the lemmas we just proved *) +Ltac do_det_step := + match goal with + | |- nsteps det_step _ _ _ => econstructor 2; first do_det_step + | |- det_step _ _ => simpl; solve[eauto 10] + end. diff --git a/theories/type_systems/systemf_mu/tactics.v b/theories/type_systems/systemf_mu/tactics.v new file mode 100644 index 0000000..70b9da1 --- /dev/null +++ b/theories/type_systems/systemf_mu/tactics.v @@ -0,0 +1,84 @@ +From stdpp Require Import gmap base relations. +From iris Require Import prelude. +From semantics.lib Require Export facts maps sets. +From semantics.ts.systemf_mu Require Export lang notation types. + +Ltac map_solver := + repeat match goal with + | |- (⤉ _) !! _ = Some _ => + rewrite fmap_insert + | |- <[ ?p := _ ]> _ !! ?p = Some _ => + apply lookup_insert + | |- <[ _ := _ ]> _ !! _ = Some _ => + rewrite lookup_insert_ne; [ | congruence] + end. +Ltac solve_typing := + intros; + repeat match goal with + | |- TY _ ; _ ⊢ ?e : ?A => first [eassumption | econstructor] + (* heuristic to solve tapp goals where we need to pick the right type for the substitution *) + | |- TY _ ; _ ⊢ ?e <> : ?A => eapply typed_tapp'; [solve_typing | | asimpl; reflexivity] + | |- TY _ ; _ ⊢ Unroll ?e : ?A => eapply typed_unroll'; [solve_typing | asimpl; reflexivity] + | |- bin_op_typed _ _ _ _ => econstructor + | |- un_op_typed _ _ _ _ => econstructor + | |- type_wf _ ?e => assert_fails (is_evar e); eassumption + | |- type_wf _ ?e => assert_fails (is_evar e); econstructor + | |- type_wf _ (subst _ ?A) => eapply type_wf_subst; [ | intros; simpl] + | |- type_wf _ ?e => assert_fails (is_evar e); eapply type_wf_mono; [eassumption | lia] + (* conditions spawned by the tyvar case of [type_wf] *) + | |- _ < _ => lia + (* conditions spawned by the variable case *) + | |- _ !! _ = Some _ => map_solver + end. + +Tactic Notation "unify_type" uconstr(A) := + match goal with + | |- TY ?n; ?Γ ⊢ ?e : ?B => unify A B + end. +Tactic Notation "replace_type" uconstr(A) := + match goal with + | |- TY ?n; ?Γ ⊢ ?e : ?B => replace B%ty with A%ty; cycle -1; first try by asimpl + end. + + +Ltac simplify_list_elem := + simpl; + repeat match goal with + | |- ?x ∈ ?y :: ?l => apply elem_of_cons; first [left; reflexivity | right] + | |- _ ∉ [] => apply not_elem_of_nil + | |- _ ∉ _ :: _ => apply not_elem_of_cons; split + end; try fast_done. +Ltac simplify_list_subseteq := + simpl; + repeat match goal with + | |- ?x :: _ ⊆ ?x :: _ => apply list_subseteq_cons_l + | |- ?x :: _ ⊆ _ => apply list_subseteq_cons_elem; first solve [simplify_list_elem] + | |- elements _ ⊆ elements _ => apply elements_subseteq; set_solver + | |- [] ⊆ _ => apply list_subseteq_nil + | |- ?x :b: _ ⊆ ?x :b: _ => apply list_subseteq_cons_binder + (* NOTE: this might make the goal unprovable *) + (*| |- elements _ ⊆ _ :: _ => apply list_subseteq_cons_r*) + end; + try fast_done. + +(* Try to solve [is_closed] goals using a number of heuristics (that shouldn't make the goal unprovable) *) +Ltac simplify_closed := + simpl; intros; + repeat match goal with + | |- closed _ _ => unfold closed; simpl + | |- Is_true (is_closed [] _) => first [assumption | done] + | |- Is_true (is_closed _ (lang.subst _ _ _)) => rewrite subst_is_closed_nil; last solve[simplify_closed] + | |- Is_true (is_closed ?X ?v) => assert_fails (is_evar X); eapply is_closed_weaken + | |- Is_true (is_closed _ _) => eassumption + | |- Is_true (_ && true) => rewrite andb_true_r + | |- Is_true (true && _) => rewrite andb_true_l + | |- Is_true (?a && ?a) => rewrite andb_diag + | |- Is_true (_ && _) => simpl; rewrite !andb_True; split_and! + | |- _ ⊆ ?A => match type of A with | list _ => simplify_list_subseteq end + | |- _ ∈ ?A => match type of A with | list _ => simplify_list_elem end + | |- _ ≠ _ => congruence + | H : closed _ _ |- _ => unfold closed in H; simpl in H + | H: Is_true (_ && _) |- _ => simpl in H; apply andb_True in H + | H : _ ∧ _ |- _ => destruct H + | |- Is_true (bool_decide (_ ∈ _)) => apply bool_decide_pack; set_solver + end; try fast_done. diff --git a/theories/type_systems/systemf_mu/types.v b/theories/type_systems/systemf_mu/types.v new file mode 100644 index 0000000..46c6842 --- /dev/null +++ b/theories/type_systems/systemf_mu/types.v @@ -0,0 +1,1073 @@ +From stdpp Require Import base relations. +From iris Require Import prelude. +From semantics.lib Require Import maps. +From semantics.ts.systemf_mu Require Import lang notation. +From Autosubst Require Export Autosubst. + +(** ** Syntactic typing *) +(** We use De Bruijn indices with the help of the Autosubst library. *) +Inductive type : Type := + (** [var] is the type of variables of Autosubst -- it unfolds to [nat] *) + | TVar : var → type + | Int + | Bool + | Unit + (** The [{bind 1 of type}] tells Autosubst to put a De Bruijn binder here *) + | TForall : {bind 1 of type} → type + | TExists : {bind 1 of type} → type + | Fun (A B : type) + | Prod (A B : type) + | Sum (A B : type) + | TMu : {bind 1 of type} → type +. + +(** Autosubst instances. + This lets Autosubst do its magic and derive all the substitution functions, etc. + *) +#[export] Instance Ids_type : Ids type. derive. Defined. +#[export] Instance Rename_type : Rename type. derive. Defined. +#[export] Instance Subst_type : Subst type. derive. Defined. +#[export] Instance SubstLemmas_typer : SubstLemmas type. derive. Qed. + +Definition typing_context := gmap string type. +Implicit Types + (Γ : typing_context) + (v : val) + (e : expr). + +Declare Scope FType_scope. +Delimit Scope FType_scope with ty. +Bind Scope FType_scope with type. +Notation "# x" := (TVar x) : FType_scope. +Infix "→" := Fun : FType_scope. +Notation "(→)" := Fun (only parsing) : FType_scope. +Notation "∀: τ" := + (TForall τ%ty) + (at level 100, τ at level 200) : FType_scope. +Notation "∃: τ" := + (TExists τ%ty) + (at level 100, τ at level 200) : FType_scope. +Infix "×" := Prod (at level 70) : FType_scope. +Notation "(×)" := Prod (only parsing) : FType_scope. +Infix "+" := Sum : FType_scope. +Notation "(+)" := Sum (only parsing) : FType_scope. +Notation "μ: A" := + (TMu A%ty) + (at level 100, A at level 200) : FType_scope. + +(** Shift all the indices in the context by one, + used when inserting a new type interpretation in Δ. *) +(* [<$>] is notation for the [fmap] operation that maps the substitution over the whole map. *) +(* [ren] is Autosubst's renaming operation -- it renames all type variables according to the given function, + in this case [(+1)] to shift the variables up by 1. *) +Notation "⤉ Γ" := (Autosubst_Classes.subst (ren (+1)) <$> Γ) (at level 10, format "⤉ Γ"). + +Reserved Notation "'TY' n ; Γ ⊢ e : A" (at level 74, e, A at next level). + +(** [type_wf n A] states that a type [A] has only free variables up to < [n]. + (in other words, all variables occurring free are strictly bounded by [n]). *) +Inductive type_wf : nat → type → Prop := + | type_wf_TVar m n: + m < n → + type_wf n (TVar m) + | type_wf_Int n: type_wf n Int + | type_wf_Bool n : type_wf n Bool + | type_wf_Unit n : type_wf n Unit + | type_wf_TForall n A : + type_wf (S n) A → + type_wf n (TForall A) + | type_wf_TExists n A : + type_wf (S n) A → + type_wf n (TExists A) + | type_wf_Fun n A B: + type_wf n A → + type_wf n B → + type_wf n (Fun A B) + | type_wf_Prod n A B : + type_wf n A → + type_wf n B → + type_wf n (Prod A B) + | type_wf_Sum n A B : + type_wf n A → + type_wf n B → + type_wf n (Sum A B) + | type_wf_mu n A : + type_wf (S n) A → + type_wf n (μ: A) +. +#[export] Hint Constructors type_wf : core. + +Inductive bin_op_typed : bin_op → type → type → type → Prop := + | plus_op_typed : bin_op_typed PlusOp Int Int Int + | minus_op_typed : bin_op_typed MinusOp Int Int Int + | mul_op_typed : bin_op_typed MultOp Int Int Int + | lt_op_typed : bin_op_typed LtOp Int Int Bool + | le_op_typed : bin_op_typed LeOp Int Int Bool + | eq_op_typed : bin_op_typed EqOp Int Int Bool. +#[export] Hint Constructors bin_op_typed : core. + +Inductive un_op_typed : un_op → type → type → Prop := + | neg_op_typed : un_op_typed NegOp Bool Bool + | minus_un_op_typed : un_op_typed MinusUnOp Int Int. + +Inductive syn_typed : nat → typing_context → expr → type → Prop := + | typed_var n Γ x A : + Γ !! x = Some A → + TY n; Γ ⊢ (Var x) : A + | typed_lam n Γ x e A B : + TY n ; (<[ x := A]> Γ) ⊢ e : B → + type_wf n A → + TY n; Γ ⊢ (Lam (BNamed x) e) : (A → B) + | typed_lam_anon n Γ e A B : + TY n ; Γ ⊢ e : B → + type_wf n A → + TY n; Γ ⊢ (Lam BAnon e) : (A → B) + | typed_tlam n Γ e A : + (* we need to shift the context up as we descend under a binder *) + TY S n; (⤉ Γ) ⊢ e : A → + TY n; Γ ⊢ (Λ, e) : (∀: A) + | typed_tapp n Γ A B e : + TY n; Γ ⊢ e : (∀: A) → + type_wf n B → + (* A.[B/] is the notation for Autosubst's substitution operation that + replaces variable 0 by [B] *) + TY n; Γ ⊢ (e <>) : (A.[B/]) + | typed_pack n Γ A B e : + type_wf n B → + type_wf (S n) A → + TY n; Γ ⊢ e : (A.[B/]) → + TY n; Γ ⊢ (pack e) : (∃: A) + | typed_unpack n Γ A B e e' x : + type_wf n B → (* we should not leak the existential! *) + TY n; Γ ⊢ e : (∃: A) → + (* As we descend under a type variable binder for the typing of [e'], + we need to shift the indices in [Γ] and [B] up by one. + On the other hand, [A] is already defined under this binder, so we need not shift it. + *) + TY (S n); (<[x := A]>(⤉Γ)) ⊢ e' : (B.[ren (+1)]) → + TY n; Γ ⊢ (unpack e as BNamed x in e') : B + | typed_int n Γ z : TY n; Γ ⊢ (Lit $ LitInt z) : Int + | typed_bool n Γ b : TY n; Γ ⊢ (Lit $ LitBool b) : Bool + | typed_unit n Γ : TY n; Γ ⊢ (Lit $ LitUnit) : Unit + | typed_if n Γ e0 e1 e2 A : + TY n; Γ ⊢ e0 : Bool → + TY n; Γ ⊢ e1 : A → + TY n; Γ ⊢ e2 : A → + TY n; Γ ⊢ If e0 e1 e2 : A + | typed_app n Γ e1 e2 A B : + TY n; Γ ⊢ e1 : (A → B) → + TY n; Γ ⊢ e2 : A → + TY n; Γ ⊢ (e1 e2)%E : B + | typed_binop n Γ e1 e2 op A B C : + bin_op_typed op A B C → + TY n; Γ ⊢ e1 : A → + TY n; Γ ⊢ e2 : B → + TY n; Γ ⊢ BinOp op e1 e2 : C + | typed_unop n Γ e op A B : + un_op_typed op A B → + TY n; Γ ⊢ e : A → + TY n; Γ ⊢ UnOp op e : B + | typed_pair n Γ e1 e2 A B : + TY n; Γ ⊢ e1 : A → + TY n; Γ ⊢ e2 : B → + TY n; Γ ⊢ (e1, e2) : A × B + | typed_fst n Γ e A B : + TY n; Γ ⊢ e : A × B → + TY n; Γ ⊢ Fst e : A + | typed_snd n Γ e A B : + TY n; Γ ⊢ e : A × B → + TY n; Γ ⊢ Snd e : B + | typed_injl n Γ e A B : + type_wf n B → + TY n; Γ ⊢ e : A → + TY n; Γ ⊢ InjL e : A + B + | typed_injr n Γ e A B : + type_wf n A → + TY n; Γ ⊢ e : B → + TY n; Γ ⊢ InjR e : A + B + | typed_case n Γ e e1 e2 A B C : + TY n; Γ ⊢ e : B + C → + TY n; Γ ⊢ e1 : (B → A) → + TY n; Γ ⊢ e2 : (C → A) → + TY n; Γ ⊢ Case e e1 e2 : A + | typed_roll n Γ e A : + TY n; Γ ⊢ e : (A.[(μ: A)/]) → + TY n; Γ ⊢ (roll e) : (μ: A) + | typed_unroll n Γ e A : + TY n; Γ ⊢ e : (μ: A) → + TY n; Γ ⊢ (unroll e) : (A.[(μ: A)/]) +where "'TY' n ; Γ ⊢ e : A" := (syn_typed n Γ e%E A%ty). +#[export] Hint Constructors syn_typed : core. + +(** Examples *) +Goal TY 0; ∅ ⊢ (λ: "x", #1 + "x")%E : (Int → Int). +Proof. eauto. Qed. +(** [∀: #0 → #0] corresponds to [∀ α. α → α] with named binders. *) +Goal TY 0; ∅ ⊢ (Λ, λ: "x", "x")%E : (∀: #0 → #0). +Proof. repeat econstructor. Qed. +Goal TY 0; ∅ ⊢ (pack ((λ: "x", "x"), #42)) : ∃: (#0 → #0) × #0. +Proof. + apply (typed_pack _ _ _ Int). + - eauto. + - repeat econstructor. + - (* [asimpl] is Autosubst's tactic for simplifying goals involving type substitutions. *) + asimpl. eauto. +Qed. +Goal TY 0; ∅ ⊢ (unpack (pack ((λ: "x", "x"), #42)) as "y" in (λ: "x", #1337) ((Fst "y") (Snd "y"))) : Int. +Proof. + (* if we want to typecheck stuff with existentials, we need a bit more explicit proofs. + Letting eauto try to instantiate the evars becomes too expensive. *) + apply (typed_unpack _ _ ((#0 → #0) × #0)%ty). + - done. + - apply (typed_pack _ _ _ Int); asimpl; eauto. + repeat econstructor. + - eapply (typed_app _ _ _ _ (#0)%ty); eauto 10. +Qed. + +(** fails: we are not allowed to leak the existential *) +Goal TY 0; ∅ ⊢ (unpack (pack ((λ: "x", "x"), #42)) as "y" in (Fst "y") (Snd "y")) : #0. +Proof. + apply (typed_unpack _ _ ((#0 → #0) × #0)%ty). +Abort. + +(* derived typing rule for match *) +Lemma typed_match n Γ e e1 e2 x1 x2 A B C : + type_wf n B → + type_wf n C → + TY n; Γ ⊢ e : B + C → + TY n; <[x1 := B]> Γ ⊢ e1 : A → + TY n; <[x2 := C]> Γ ⊢ e2 : A → + TY n; Γ ⊢ match: e with InjL (BNamed x1) => e1 | InjR (BNamed x2) => e2 end : A. +Proof. eauto. Qed. + +Lemma syn_typed_closed n Γ e A X : + TY n ; Γ ⊢ e : A → + (∀ x, x ∈ dom Γ → x ∈ X) → + is_closed X e. +Proof. + induction 1 as [ | ??????? IH | | n Γ e A H IH | | | n Γ A B e e' x Hwf H1 IH1 H2 IH2 | | | | | | | | | | | | | | | ] in X |-*; simpl; intros Hx; try done. + + { (* var *) apply bool_decide_pack, Hx. apply elem_of_dom; eauto. } + { (* lam *) apply IH. + intros y. rewrite elem_of_dom lookup_insert_is_Some. + intros [<- | [? Hy]]; first by apply elem_of_cons; eauto. + apply elem_of_cons. right. eapply Hx. by apply elem_of_dom. + } + { (* anon lam *) naive_solver. } + { (* tlam *) + eapply IH. intros x Hel. apply Hx. + by rewrite dom_fmap in Hel. + } + 3: { (* unpack *) + apply andb_True; split. + - apply IH1. apply Hx. + - apply IH2. intros y. rewrite elem_of_dom lookup_insert_is_Some. + intros [<- | [? Hy]]; first by apply elem_of_cons; eauto. + apply elem_of_cons. right. eapply Hx. + apply elem_of_dom. revert Hy. rewrite lookup_fmap fmap_is_Some. done. + } + (* everything else *) + all: repeat match goal with + | |- Is_true (_ && _) => apply andb_True; split + end. + all: try naive_solver. +Qed. + +(** *** Lemmas about [type_wf] *) +Lemma type_wf_mono n m A: + type_wf n A → n ≤ m → type_wf m A. +Proof. + induction 1 in m |-*; eauto with lia. +Qed. + +Lemma type_wf_rename n A δ: + type_wf n A → + (∀ i j, i < j → δ i < δ j) → + type_wf (δ n) (rename δ A). +Proof. + induction 1 in δ |-*; intros Hmon; simpl; eauto. + all: econstructor; eapply type_wf_mono; first eapply IHtype_wf; last done. + all: intros i j Hlt; destruct i, j; simpl; try lia. + all: rewrite -Nat.succ_lt_mono; eapply Hmon; lia. +Qed. + +(** [A.[σ]], i.e. [A] with the substitution [σ] applied to it, is well-formed under [m] if + [A] is well-formed under [n] and all the things we substitute up to [n] are well-formed under [m]. + *) +Lemma type_wf_subst n m A σ: + type_wf n A → + (∀ x, x < n → type_wf m (σ x)) → + type_wf m A.[σ]. +Proof. + induction 1 in m, σ |-*; intros Hsub; simpl; eauto. + + econstructor; eapply IHtype_wf. + intros [|x]; rewrite /up //=. + - econstructor. lia. + - intros Hlt % Nat.succ_lt_mono. eapply type_wf_rename; eauto. + intros i j Hlt'; simpl; lia. + + econstructor; eapply IHtype_wf. + intros [|x]; rewrite /up //=. + - econstructor. lia. + - intros Hlt % Nat.succ_lt_mono. eapply type_wf_rename; eauto. + intros i j Hlt'; simpl; lia. + + econstructor. eapply IHtype_wf. + intros [|x]; rewrite /up //=. + - econstructor. lia. + - intros Hlt % Nat.succ_lt_mono. eapply type_wf_rename; eauto. + intros ???. simpl; lia. +Qed. + +Fixpoint free_vars A : nat → Prop := + match A with + | TVar n => λ m, m = n + | Int => λ _, False + | Bool => λ _, False + | Unit => λ _, False + | Fun A B => λ n, free_vars A n ∨ free_vars B n + | Prod A B => λ n, free_vars A n ∨ free_vars B n + | Sum A B => λ n, free_vars A n ∨ free_vars B n + | TForall A => λ n, free_vars A (S n) + | TExists A => λ n, free_vars A (S n) + | TMu A => λ n, free_vars A (S n) + end. + +Definition bounded n A := + (∀ x, free_vars A x → x < n). + +Lemma type_wf_bounded n A: + type_wf n A ↔ bounded n A. +Proof. + rewrite /bounded; split. + - induction 1; simpl; try naive_solver. + + intros x Hfree % IHtype_wf. lia. + + intros x Hfree % IHtype_wf. lia. + + intros x Hfree % IHtype_wf. lia. + - induction A in n |-*; simpl; eauto. + + intros Hsub. econstructor. eapply IHA. + intros ??. destruct x as [|x]; first lia. + eapply Hsub in H. lia. + + intros Hsub. econstructor. eapply IHA. + intros ??. destruct x as [|x]; first lia. + eapply Hsub in H. lia. + + intros Hsub. econstructor; eauto. + + intros Hsub. econstructor; eauto. + + intros Hsub. econstructor; eauto. + + intros Hsub. econstructor. eapply IHA. + intros ??. destruct x as [|x]; first lia. + eapply Hsub in H. lia. +Qed. + +Lemma free_vars_rename A x δ: + free_vars A x → free_vars (rename δ A) (δ x). +Proof. + induction A in x, δ |-*; simpl; try naive_solver. + - intros Hf. apply (IHA (S x) (upren δ) Hf). + - intros Hf. apply (IHA (S x) (upren δ) Hf). + - intros Hf. apply (IHA (S x) (upren δ) Hf). +Qed. + +Lemma free_vars_subst x n A σ : + bounded n A.[σ] → free_vars A x → bounded n (σ x). +Proof. + induction A in n, σ, x |-*; simpl; try naive_solver. + - rewrite -type_wf_bounded. inversion 1; subst. revert H2; clear H. + rewrite type_wf_bounded. + intros Hbd Hfree. eapply IHA in Hbd; last done. + revert Hbd. rewrite /up //=. + intros Hbd y Hf. enough (S y < S n) by lia. + eapply Hbd. simpl. by eapply free_vars_rename. + - rewrite -type_wf_bounded. inversion 1; subst. revert H2; clear H. + rewrite type_wf_bounded. + intros Hbd Hfree. eapply IHA in Hbd; last done. + revert Hbd. rewrite /up //=. + intros Hbd y Hf. enough (S y < S n) by lia. + eapply Hbd. simpl. by eapply free_vars_rename. + - rewrite -!type_wf_bounded. inversion 1; subst. + revert H3 H4. rewrite !type_wf_bounded. naive_solver. + - rewrite -!type_wf_bounded. inversion 1; subst. + revert H3 H4. rewrite !type_wf_bounded. naive_solver. + - rewrite -!type_wf_bounded. inversion 1; subst. + revert H3 H4. rewrite !type_wf_bounded. naive_solver. + - rewrite -type_wf_bounded. inversion 1; subst. revert H2; clear H. + rewrite type_wf_bounded. + intros Hbd Hfree. eapply IHA in Hbd; last done. + revert Hbd. rewrite /up //=. + intros Hbd y Hf. enough (S y < S n) by lia. + eapply Hbd. simpl. by eapply free_vars_rename. +Qed. + +Lemma type_wf_rec_type n A: + type_wf n A.[(μ: A)%ty/] → type_wf (S n) A. +Proof. + rewrite !type_wf_bounded. intros Hbound x Hfree. + eapply free_vars_subst in Hbound; last done. + destruct x as [|x]; first lia; simpl in Hbound. + eapply type_wf_bounded in Hbound. inversion Hbound; subst; lia. +Qed. + +Lemma type_wf_single_subst n A B: type_wf n B → type_wf (S n) A → type_wf n A.[B/]. +Proof. + intros HB HA. eapply type_wf_subst; first done. + intros [|x]; simpl; eauto. + intros ?; econstructor. lia. +Qed. + +(** We lift [type_wf] to well-formedness of contexts *) +Definition ctx_wf n Γ := (∀ x A, Γ !! x = Some A → type_wf n A). + +Lemma ctx_wf_empty n : ctx_wf n ∅. +Proof. rewrite /ctx_wf. set_solver. Qed. + +Lemma ctx_wf_insert n x Γ A: ctx_wf n Γ → type_wf n A → ctx_wf n (<[x := A]> Γ). +Proof. intros H1 H2 y B. rewrite lookup_insert_Some. naive_solver. Qed. + +Lemma ctx_wf_up n Γ: + ctx_wf n Γ → ctx_wf (S n) (⤉Γ). +Proof. + intros Hwf x A; rewrite lookup_fmap. + intros (B & Hlook & ->) % fmap_Some. + asimpl. eapply type_wf_subst; first eauto. + intros y Hlt. simpl. econstructor. lia. +Qed. + +#[global] +Hint Resolve ctx_wf_empty ctx_wf_insert ctx_wf_up : core. + +(** Well-typed terms at [A] under a well-formed context have well-formed types [A].*) +Lemma syn_typed_wf n Γ e A: + ctx_wf n Γ → + TY n; Γ ⊢ e : A → + type_wf n A. +Proof. + intros Hwf; induction 1 as [ | n Γ x e A B Hty IH Hwfty | | n Γ e A Hty IH | n Γ A B e Hty IH Hwfty | n Γ A B e Hwfty Hty IH| | | | | | n Γ e1 e2 A B HtyA IHA HtyB IHB | n Γ e1 e2 op A B C Hop HtyA IHA HtyB IHB | n Γ e op A B Hop H IH | n Γ e1 e2 A B HtyA IHA HtyB IHB | n Γ e A B Hty IH | n Γ e A B Hty IH | n Γ e A B Hwfty Hty IH | n Γ e A B Hwfty Hty IH| n Γ e e1 e2 A B C Htye IHe Htye1 IHe1 Htye2 IHe2 | n Γ e A Hty IH | n Γ e A Hty IH]; eauto. + - eapply type_wf_single_subst; first done. + specialize (IH Hwf) as Hwf'. + by inversion Hwf'. + - specialize (IHA Hwf) as Hwf'. + by inversion Hwf'; subst. + - inversion Hop; subst; eauto. + - inversion Hop; subst; eauto. + - specialize (IH Hwf) as Hwf'. by inversion Hwf'; subst. + - specialize (IH Hwf) as Hwf'. by inversion Hwf'; subst. + - specialize (IHe1 Hwf) as Hwf''. by inversion Hwf''; subst. + - specialize (IH Hwf) as Hwf'%type_wf_rec_type. + by econstructor. + - eapply type_wf_single_subst; first by apply IH. + specialize (IH Hwf) as Hwf'. + by inversion Hwf'. +Qed. + +Lemma renaming_inclusion Γ Δ : Γ ⊆ Δ → ⤉Γ ⊆ ⤉Δ. +Proof. + eapply map_fmap_mono. +Qed. + +Lemma typed_weakening n m Γ Δ e A: + TY n; Γ ⊢ e : A → + Γ ⊆ Δ → + n ≤ m → + TY m; Δ ⊢ e : A. +Proof. + induction 1 as [| n Γ x e A B Htyp IH | | n Γ e A Htyp IH | | |n Γ A B e e' x Hwf H1 IH1 H2 IH2 | | | | | | | | | | | | | | | ] in Δ, m |-*; intros Hsub Hle; eauto using type_wf_mono. + - (* var *) econstructor. by eapply lookup_weaken. + - (* lam *) econstructor; last by eapply type_wf_mono. eapply IH; eauto. by eapply insert_mono. + - (* tlam *) econstructor. eapply IH; last by lia. by eapply renaming_inclusion. + - (* pack *) + econstructor; last naive_solver. all: (eapply type_wf_mono; [ done | lia]). + - (* unpack *) econstructor. + + eapply type_wf_mono; done. + + eapply IH1; done. + + eapply IH2; last lia. apply insert_mono. by apply renaming_inclusion. +Qed. + +Lemma type_wf_subst_dom σ τ n A: + type_wf n A → + (∀ m, m < n → σ m = τ m) → + A.[σ] = A.[τ]. +Proof. + induction 1 in σ, τ |-*; simpl; eauto. + - (* tforall *) + intros Heq; asimpl. f_equal. + eapply IHtype_wf; intros [|m]; rewrite /up; simpl; first done. + intros Hlt. f_equal. eapply Heq. lia. + - (* texists *) + intros Heq; asimpl. f_equal. + eapply IHtype_wf. intros [ | m]; rewrite /up; simpl; first done. + intros Hlt. f_equal. apply Heq. lia. + - (* fun *) intros ?. f_equal; eauto. + - (* prod *) intros ?. f_equal; eauto. + - (* sum *) intros ?. f_equal; eauto. + - (* rec *) + intros Heq; asimpl. f_equal. + eapply IHtype_wf; intros [|m]; rewrite /up; simpl; first done. + intros Hlt. f_equal. eapply Heq. lia. +Qed. + +Lemma type_wf_closed A σ: + type_wf 0 A → + A.[σ] = A. +Proof. + intros Hwf; erewrite (type_wf_subst_dom _ (ids) 0). + - by asimpl. + - done. + - intros ??; lia. +Qed. + +(** Typing inversion lemmas *) +Lemma var_inversion Γ n (x: string) A: TY n; Γ ⊢ x : A → Γ !! x = Some A. +Proof. inversion 1; subst; auto. Qed. + +Lemma lam_inversion n Γ (x: string) e C: + TY n; Γ ⊢ (λ: x, e) : C → + ∃ A B, C = (A → B)%ty ∧ type_wf n A ∧ TY n; <[x:=A]> Γ ⊢ e : B. +Proof. inversion 1; subst; eauto 10. Qed. + +Lemma lam_anon_inversion n Γ e C: + TY n; Γ ⊢ (λ: <>, e) : C → + ∃ A B, C = (A → B)%ty ∧ type_wf n A ∧ TY n; Γ ⊢ e : B. +Proof. inversion 1; subst; eauto 10. Qed. + +Lemma app_inversion n Γ e1 e2 B: + TY n; Γ ⊢ e1 e2 : B → + ∃ A, TY n; Γ ⊢ e1 : (A → B) ∧ TY n; Γ ⊢ e2 : A. +Proof. inversion 1; subst; eauto. Qed. + +Lemma if_inversion n Γ e0 e1 e2 B: + TY n; Γ ⊢ If e0 e1 e2 : B → + TY n; Γ ⊢ e0 : Bool ∧ TY n; Γ ⊢ e1 : B ∧ TY n; Γ ⊢ e2 : B. +Proof. inversion 1; subst; eauto. Qed. + +Lemma binop_inversion n Γ op e1 e2 B: + TY n; Γ ⊢ BinOp op e1 e2 : B → + ∃ A1 A2, bin_op_typed op A1 A2 B ∧ TY n; Γ ⊢ e1 : A1 ∧ TY n; Γ ⊢ e2 : A2. +Proof. inversion 1; subst; eauto. Qed. + +Lemma unop_inversion n Γ op e B: + TY n; Γ ⊢ UnOp op e : B → + ∃ A, un_op_typed op A B ∧ TY n; Γ ⊢ e : A. +Proof. inversion 1; subst; eauto. Qed. + +Lemma type_app_inversion n Γ e B: + TY n; Γ ⊢ e <> : B → + ∃ A C, B = A.[C/] ∧ type_wf n C ∧ TY n; Γ ⊢ e : (∀: A). +Proof. inversion 1; subst; eauto. Qed. + +Lemma type_lam_inversion n Γ e B: + TY n; Γ ⊢ (Λ,e) : B → + ∃ A, B = (∀: A)%ty ∧ TY (S n); ⤉Γ ⊢ e : A. +Proof. inversion 1; subst; eauto. Qed. + +Lemma type_pack_inversion n Γ e B : + TY n; Γ ⊢ (pack e) : B → + ∃ A C, B = (∃: A)%ty ∧ TY n; Γ ⊢ e : (A.[C/])%ty ∧ type_wf n C ∧ type_wf (S n) A. +Proof. inversion 1; subst; eauto 10. Qed. + +Lemma type_unpack_inversion n Γ e e' x B : + TY n; Γ ⊢ (unpack e as x in e') : B → + ∃ A x', x = BNamed x' ∧ type_wf n B ∧ TY n; Γ ⊢ e : (∃: A) ∧ TY S n; <[x' := A]> (⤉Γ) ⊢ e' : (B.[ren (+1)]). +Proof. inversion 1; subst; eauto 10. Qed. + +Lemma pair_inversion n Γ e1 e2 C : + TY n; Γ ⊢ (e1, e2) : C → + ∃ A B, C = (A × B)%ty ∧ TY n; Γ ⊢ e1 : A ∧ TY n; Γ ⊢ e2 : B. +Proof. inversion 1; subst; eauto. Qed. + +Lemma fst_inversion n Γ e A : + TY n; Γ ⊢ Fst e : A → + ∃ B, TY n; Γ ⊢ e : A × B. +Proof. inversion 1; subst; eauto. Qed. + +Lemma snd_inversion n Γ e B : + TY n; Γ ⊢ Snd e : B → + ∃ A, TY n; Γ ⊢ e : A × B. +Proof. inversion 1; subst; eauto. Qed. + +Lemma injl_inversion n Γ e C : + TY n; Γ ⊢ InjL e : C → + ∃ A B, C = (A + B)%ty ∧ TY n; Γ ⊢ e : A ∧ type_wf n B. +Proof. inversion 1; subst; eauto. Qed. + +Lemma injr_inversion n Γ e C : + TY n; Γ ⊢ InjR e : C → + ∃ A B, C = (A + B)%ty ∧ TY n; Γ ⊢ e : B ∧ type_wf n A. +Proof. inversion 1; subst; eauto. Qed. + +Lemma case_inversion n Γ e e1 e2 A : + TY n; Γ ⊢ Case e e1 e2 : A → + ∃ B C, TY n; Γ ⊢ e : B + C ∧ TY n; Γ ⊢ e1 : (B → A) ∧ TY n; Γ ⊢ e2 : (C → A). +Proof. inversion 1; subst; eauto. Qed. + +Lemma roll_inversion n Γ e B: + TY n; Γ ⊢ (roll e) : B → + ∃ A, B = (μ: A)%ty ∧ TY n; Γ ⊢ e : A.[μ: A/]. +Proof. inversion 1; subst; eauto. Qed. + +Lemma unroll_inversion n Γ e B: + TY n; Γ ⊢ (unroll e) : B → + ∃ A, B = (A.[μ: A/])%ty ∧ TY n; Γ ⊢ e : μ: A. +Proof. inversion 1; subst; eauto. Qed. + + +Lemma typed_substitutivity n e e' Γ (x: string) A B : + TY 0; ∅ ⊢ e' : A → + TY n; (<[x := A]> Γ) ⊢ e : B → + TY n; Γ ⊢ lang.subst x e' e : B. +Proof. + intros He'. revert n B Γ; induction e as [| y | y | | | | | | | | | | | | | | | | ]; intros n B Γ; simpl. + - inversion 1; subst; auto. + - intros Hp % var_inversion. + destruct (decide (x = y)). + + subst. rewrite lookup_insert in Hp. injection Hp as ->. + eapply typed_weakening; [done| |lia]. apply map_empty_subseteq. + + rewrite lookup_insert_ne in Hp; last done. auto. + - destruct y as [ | y]. + { intros (A' & C & -> & Hwf & Hty) % lam_anon_inversion. + econstructor; last done. destruct decide as [Heq|]. + + congruence. + + eauto. + } + intros (A' & C & -> & Hwf & Hty) % lam_inversion. + econstructor; last done. destruct decide as [Heq|]. + + injection Heq as [= ->]. by rewrite insert_insert in Hty. + + rewrite insert_commute in Hty; last naive_solver. eauto. + - intros (C & Hty1 & Hty2) % app_inversion. eauto. + - intros (? & Hop & H1) % unop_inversion. + destruct op; inversion Hop; subst; eauto. + - intros (? & ? & Hop & H1 & H2) % binop_inversion. + destruct op; inversion Hop; subst; eauto. + - intros (H1 & H2 & H3)%if_inversion. naive_solver. + - intros (C & D & -> & Hwf & Hty) % type_app_inversion. eauto. + - intros (C & -> & Hty)%type_lam_inversion. econstructor. + eapply IHe. revert Hty. rewrite fmap_insert. + eapply syn_typed_wf in He'; last by naive_solver. + rewrite type_wf_closed; eauto. + - intros (C & D & -> & Hty & Hwf1 & Hwf2)%type_pack_inversion. + econstructor; [done..|]. apply IHe. done. + - intros (C & x' & -> & Hwf & Hty1 & Hty2)%type_unpack_inversion. + econstructor; first done. + + eapply IHe1. done. + + destruct decide as [Heq | ]. + * injection Heq as [= ->]. by rewrite fmap_insert insert_insert in Hty2. + * rewrite fmap_insert in Hty2. rewrite insert_commute in Hty2; last naive_solver. + eapply IHe2. rewrite type_wf_closed in Hty2; first done. + eapply syn_typed_wf; last apply He'. done. + - intros (? & ? & -> & ? & ?) % pair_inversion. eauto. + - intros (? & ?)%fst_inversion. eauto. + - intros (? & ?)%snd_inversion. eauto. + - intros (? & ? & -> & ? & ?)%injl_inversion. eauto. + - intros (? & ? & -> & ? & ?)%injr_inversion. eauto. + - intros (? & ? & ? & ? & ?)%case_inversion. eauto. + - intros (C & -> & Hty) % roll_inversion. eauto. + - intros (C & -> & Hty) % unroll_inversion. eauto. +Qed. + +(** Canonical values *) +Lemma canonical_values_arr n Γ e A B: + TY n; Γ ⊢ e : (A → B) → + is_val e → + ∃ x e', e = (λ: x, e')%E. +Proof. + inversion 1; simpl; naive_solver. +Qed. + +Lemma canonical_values_forall n Γ e A: + TY n; Γ ⊢ e : (∀: A)%ty → + is_val e → + ∃ e', e = (Λ, e')%E. +Proof. + inversion 1; simpl; naive_solver. +Qed. + +Lemma canonical_values_exists n Γ e A : + TY n; Γ ⊢ e : (∃: A) → + is_val e → + ∃ e', e = (pack e')%E. +Proof. inversion 1; simpl; naive_solver. Qed. + +Lemma canonical_values_int n Γ e: + TY n; Γ ⊢ e : Int → + is_val e → + ∃ n: Z, e = (#n)%E. +Proof. + inversion 1; simpl; naive_solver. +Qed. + +Lemma canonical_values_bool n Γ e: + TY n; Γ ⊢ e : Bool → + is_val e → + ∃ b: bool, e = (#b)%E. +Proof. + inversion 1; simpl; naive_solver. +Qed. + +Lemma canonical_values_unit n Γ e: + TY n; Γ ⊢ e : Unit → + is_val e → + e = (#LitUnit)%E. +Proof. + inversion 1; simpl; naive_solver. +Qed. + +Lemma canonical_values_prod n Γ e A B : + TY n; Γ ⊢ e : A × B → + is_val e → + ∃ e1 e2, e = (e1, e2)%E ∧ is_val e1 ∧ is_val e2. +Proof. + inversion 1; simpl; naive_solver. +Qed. + +Lemma canonical_values_sum n Γ e A B : + TY n; Γ ⊢ e : A + B → + is_val e → + (∃ e', e = InjL e' ∧ is_val e') ∨ (∃ e', e = InjR e' ∧ is_val e'). +Proof. + inversion 1; simpl; naive_solver. +Qed. + +Lemma canonical_values_rec n Γ e A: + TY n; Γ ⊢ e : (μ: A) → + is_val e → + ∃ e', e = (roll e')%E ∧ is_val e'. +Proof. + inversion 1; simpl; subst; naive_solver. +Qed. + +(** Progress *) +Lemma typed_progress e A: + TY 0; ∅ ⊢ e : A → is_val e ∨ reducible e. +Proof. + remember ∅ as Γ. remember 0 as n. + induction 1 as [| | | | n Γ A B e Hty IH | n Γ A B e Hwf Hwf' Hty IH | n Γ A B e e' x Hwf Hty1 IH1 Hty2 IH2 | | | | n Γ e0 e1 e2 A Hty1 IH1 Hty2 IH2 Hty3 IH3 | n Γ e1 e2 A B Hty IH1 _ IH2 | n Γ e1 e2 op A B C Hop Hty1 IH1 Hty2 IH2 | n Γ e op A B Hop Hty IH | n Γ e1 e2 A B Hty1 IH1 Hty2 IH2 | n Γ e A B Hty IH | n Γ e A B Hty IH | n Γ e A B Hwf Hty IH | n Γ e A B Hwf Hty IH| n Γ e e1 e2 A B C Htye IHe Htye1 IHe1 Htye2 IHe2 | n Γ e A Hty IH | n Γ e A Hty IH]. + - subst. naive_solver. + - left. done. + - left. done. + - left; done. + - right. destruct (IH HeqΓ Heqn) as [H1|H1]. + + eapply canonical_values_forall in Hty as [e' ->]; last done. + eexists. eapply base_contextual_step. eapply TBetaS. + + destruct H1 as [e' H1]. eexists. + eapply (fill_contextual_step [TAppCtx]). done. + - (* pack *) + destruct (IH HeqΓ Heqn) as [H | H]. + + by left. + + right. destruct H as [e' H]. eexists. eapply (fill_contextual_step [PackCtx]). done. + - (* unpack *) + destruct (IH1 HeqΓ Heqn) as [H | H]. + + eapply canonical_values_exists in Hty1 as [e'' ->]; last done. + right. eexists. eapply base_contextual_step. constructor; last done. + done. + + right. destruct H as [e'' H]. eexists. + eapply (fill_contextual_step [UnpackCtx _ _]). done. + - (* int *)left. done. + - (* bool*) left. done. + - (* unit *) left. done. + - (* if *) + destruct (IH1 HeqΓ Heqn) as [H1 | H1]. + + eapply canonical_values_bool in Hty1 as (b & ->); last done. + right. destruct b; eexists; eapply base_contextual_step; constructor. + + right. destruct H1 as [e0' Hstep]. + eexists. by eapply (fill_contextual_step [IfCtx _ _]). + - (* app *) + destruct (IH2 HeqΓ Heqn) as [H2|H2]; [destruct (IH1 HeqΓ Heqn) as [H1|H1]|]. + + eapply canonical_values_arr in Hty as (x & e & ->); last done. + right. eexists. + eapply base_contextual_step, BetaS; eauto. + + right. eapply is_val_spec in H2 as [v Heq]. + replace e2 with (of_val v); last by eapply of_to_val. + destruct H1 as [e1' Hstep]. + eexists. eapply (fill_contextual_step [AppLCtx v]). done. + + right. destruct H2 as [e2' H2]. + eexists. eapply (fill_contextual_step [AppRCtx e1]). done. + - (* binop *) + assert (A = Int ∧ B = Int) as [-> ->]. + { inversion Hop; subst A B C; done. } + destruct (IH2 HeqΓ Heqn) as [H2|H2]; [destruct (IH1 HeqΓ Heqn) as [H1|H1]|]. + + right. eapply canonical_values_int in Hty1 as [n1 ->]; last done. + eapply canonical_values_int in Hty2 as [n2 ->]; last done. + inversion Hop; subst; simpl. + all: eexists; eapply base_contextual_step; eapply BinOpS; eauto. + + right. eapply is_val_spec in H2 as [v Heq]. + replace e2 with (of_val v); last by eapply of_to_val. + destruct H1 as [e1' Hstep]. + eexists. eapply (fill_contextual_step [BinOpLCtx op v]). done. + + right. destruct H2 as [e2' H2]. + eexists. eapply (fill_contextual_step [BinOpRCtx op e1]). done. + - (* unop *) + inversion Hop; subst A B op. + + right. destruct (IH HeqΓ Heqn) as [H2 | H2]. + * eapply canonical_values_bool in Hty as [b ->]; last done. + eexists; eapply base_contextual_step; eapply UnOpS; eauto. + * destruct H2 as [e' H2]. eexists. + eapply (fill_contextual_step [UnOpCtx _]). done. + + right. destruct (IH HeqΓ Heqn) as [H2 | H2]. + * eapply canonical_values_int in Hty as [z ->]; last done. + eexists; eapply base_contextual_step; eapply UnOpS; eauto. + * destruct H2 as [e' H2]. eexists. + eapply (fill_contextual_step [UnOpCtx _]). done. + - (* pair *) + destruct (IH2 HeqΓ Heqn) as [H2|H2]; [destruct (IH1 HeqΓ Heqn) as [H1|H1]|]. + + left. done. + + right. eapply is_val_spec in H2 as [v Heq]. + replace e2 with (of_val v); last by eapply of_to_val. + destruct H1 as [e1' Hstep]. + eexists. eapply (fill_contextual_step [PairLCtx v]). done. + + right. destruct H2 as [e2' H2]. + eexists. eapply (fill_contextual_step [PairRCtx e1]). done. + - (* fst *) + destruct (IH HeqΓ Heqn) as [H | H]. + + eapply canonical_values_prod in Hty as (e1 & e2 & -> & ? & ?); last done. + right. eexists. eapply base_contextual_step. econstructor; done. + + right. destruct H as [e' H]. + eexists. eapply (fill_contextual_step [FstCtx]). done. + - (* snd *) + destruct (IH HeqΓ Heqn) as [H | H]. + + eapply canonical_values_prod in Hty as (e1 & e2 & -> & ? & ?); last done. + right. eexists. eapply base_contextual_step. econstructor; done. + + right. destruct H as [e' H]. + eexists. eapply (fill_contextual_step [SndCtx]). done. + - (* injl *) + destruct (IH HeqΓ Heqn) as [H | H]. + + left. done. + + right. destruct H as [e' H]. + eexists. eapply (fill_contextual_step [InjLCtx]). done. + - (* injr *) + destruct (IH HeqΓ Heqn) as [H | H]. + + left. done. + + right. destruct H as [e' H]. + eexists. eapply (fill_contextual_step [InjRCtx]). done. + - (* case *) + right. destruct (IHe HeqΓ Heqn) as [H1|H1]. + + eapply canonical_values_sum in Htye as [(e' & -> & ?) | (e' & -> & ?)]; last done. + * eexists. eapply base_contextual_step. econstructor. done. + * eexists. eapply base_contextual_step. econstructor. done. + + destruct H1 as [e' H1]. eexists. + eapply (fill_contextual_step [CaseCtx e1 e2]). done. + - (* roll *) + destruct (IH HeqΓ Heqn) as [Hval|Hred]. + + by left. + + right. destruct Hred as [e' Hred]. + eexists. eapply (fill_contextual_step [RollCtx]). done. + - (* unroll *) + destruct (IH HeqΓ Heqn) as [Hval|Hred]. + + eapply canonical_values_rec in Hty as (e' & -> & Hval'); last done. + right. eexists. eapply base_contextual_step. by econstructor. + + right. destruct Hred as [e' Hred]. + eexists. eapply (fill_contextual_step [UnrollCtx]). done. +Qed. + + +Definition ectx_item_typing (K: ectx_item) (A B: type) := + ∀ e, TY 0; ∅ ⊢ e : A → TY 0; ∅ ⊢ (fill_item K e) : B. + +Notation ectx := (list ectx_item). +Definition ectx_typing (K: ectx) (A B: type) := + ∀ e, TY 0; ∅ ⊢ e : A → TY 0; ∅ ⊢ (fill K e) : B. + +Lemma fill_item_typing_decompose k e A: + TY 0; ∅ ⊢ fill_item k e : A → + ∃ B, TY 0; ∅ ⊢ e : B ∧ ectx_item_typing k B A. +Proof. + unfold ectx_item_typing; destruct k; simpl; inversion 1; subst; eauto. +Qed. + +Lemma fill_item_typing_compose k e A B: + TY 0; ∅ ⊢ e : B → + ectx_item_typing k B A → + TY 0; ∅ ⊢ fill_item k e : A. +Proof. + intros H1 H2; by eapply H2. +Qed. + +Lemma fill_typing_decompose K e A: + TY 0; ∅ ⊢ fill K e : A → + ∃ B, TY 0; ∅ ⊢ e : B ∧ ectx_typing K B A. +Proof. + unfold ectx_typing; revert e; induction K as [|k K]; intros e; simpl; eauto. + intros [B [Hit Hty]] % IHK. + eapply fill_item_typing_decompose in Hit as [B' [? ?]]; eauto. +Qed. + +Lemma fill_typing_compose K e A B: + TY 0; ∅ ⊢ e : B → + ectx_typing K B A → + TY 0; ∅ ⊢ fill K e : A. +Proof. + intros H1 H2; by eapply H2. +Qed. + +Lemma fmap_up_subst σ Γ: ⤉(subst σ <$> Γ) = subst (up σ) <$> ⤉Γ. +Proof. + rewrite -!map_fmap_compose. + eapply map_fmap_ext. intros x A _. by asimpl. +Qed. + +Lemma typed_subst_type n m Γ e A σ: + TY n; Γ ⊢ e : A → (∀ k, k < n → type_wf m (σ k)) → TY m; (subst σ) <$> Γ ⊢ e : A.[σ]. +Proof. + induction 1 as [ n Γ x A Heq | | | n Γ e A Hty IH | |n Γ A B e Hwf Hwf' Hty IH | n Γ A B e e' x Hwf Hty1 IH1 Hty2 IH2 | | | | | |? ? ? ? ? ? ? ? Hop | ? ? ? ? ? ? Hop | | | | | | | | ] in σ, m |-*; simpl; intros Hlt; eauto. + - econstructor. rewrite lookup_fmap Heq //=. + - econstructor; last by eapply type_wf_subst. + rewrite -fmap_insert. eauto. + - econstructor; last by eapply type_wf_subst. eauto. + - econstructor. rewrite fmap_up_subst. eapply IH. + intros [| x] Hlt'; rewrite /up //=. + + econstructor. lia. + + eapply type_wf_rename; last by (intros ???; simpl; lia). + eapply Hlt. lia. + - replace (A.[B/].[σ]) with (A.[up σ].[B.[σ]/]) by by asimpl. + eauto using type_wf_subst. + - (* pack *) + eapply (typed_pack _ _ _ (subst σ B)). + + eapply type_wf_subst; done. + + eapply type_wf_subst; first done. + intros [ | k] Hk; first ( asimpl;constructor; lia). + rewrite /up //=. eapply type_wf_rename; last by (intros ???; simpl; lia). + eapply Hlt. lia. + + replace (A.[up σ].[B.[σ]/]) with (A.[B/].[σ]) by by asimpl. + eauto using type_wf_subst. + - (* unpack *) + eapply (typed_unpack _ _ A.[up σ]). + + eapply type_wf_subst; done. + + replace (∃: A.[up σ])%ty with ((∃: A).[σ])%ty by by asimpl. + eapply IH1. done. + + rewrite fmap_up_subst. rewrite -fmap_insert. + replace (B.[σ].[ren (+1)]) with (B.[ren(+1)].[up σ]) by by asimpl. + eapply IH2. + intros [ | k] Hk; asimpl; first (constructor; lia). + eapply type_wf_subst; first (eapply Hlt; lia). + intros k' Hk'. asimpl. constructor. lia. + - (* binop *) + inversion Hop; subst. + all: econstructor; naive_solver. + - (* unop *) + inversion Hop; subst. + all: econstructor; naive_solver. + - econstructor; last naive_solver. by eapply type_wf_subst. + - econstructor; last naive_solver. by eapply type_wf_subst. + - (* roll *) + econstructor. + replace (A.[up σ].[μ: A.[up σ]/])%ty with (A.[μ: A/].[σ])%ty by by asimpl. eauto. + - (* unroll *) + replace (A.[μ: A/].[σ])%ty with (A.[up σ].[μ: A.[up σ]/])%ty by by asimpl. + econstructor. eapply IHsyn_typed. done. +Qed. + +Lemma typed_subst_type_closed C e A: + type_wf 0 C → TY 1; ⤉∅ ⊢ e : A → TY 0; ∅ ⊢ e : A.[C/]. +Proof. + intros Hwf Hty. eapply typed_subst_type with (σ := C .: ids) (m := 0) in Hty; last first. + { intros [|k] Hlt; last lia. done. } + revert Hty. by rewrite !fmap_empty. +Qed. + +Lemma typed_subst_type_closed' x C B e A: + type_wf 0 A → + type_wf 1 C → + type_wf 0 B → + TY 1; <[x := C]> ∅ ⊢ e : A → + TY 0; <[x := C.[B/]]> ∅ ⊢ e : A. +Proof. + intros ??? Hty. + set (s := (subst (B.:ids))). + rewrite -(fmap_empty s) -(fmap_insert s). + replace A with (A.[B/]). + 2: { replace A with (A.[ids]) at 2 by by asimpl. + eapply type_wf_subst_dom; first done. lia. + } + eapply typed_subst_type; first done. + intros [ | k] Hk; last lia. done. +Qed. + +Lemma typed_preservation_base_step e e' A: + TY 0; ∅ ⊢ e : A → + base_step e e' → + TY 0; ∅ ⊢ e' : A. +Proof. + intros Hty Hstep. destruct Hstep as [ | | | op e v v' Heq Heval | op e1 v1 e2 v2 v3 Heq1 Heq2 Heval | | | | | | | ]; subst. + - eapply app_inversion in Hty as (B & H1 & H2). + destruct x as [|x]. + { eapply lam_anon_inversion in H1 as (C & D & [= -> ->] & Hwf & Hty). done. } + eapply lam_inversion in H1 as (C & D & Heq & Hwf & Hty). + injection Heq as -> ->. + eapply typed_substitutivity; eauto. + - eapply type_app_inversion in Hty as (B & C & -> & Hwf & Hty). + eapply type_lam_inversion in Hty as (A & Heq & Hty). + injection Heq as ->. by eapply typed_subst_type_closed. + - eapply type_unpack_inversion in Hty as (B & x' & -> & Hwf & Hty1 & Hty2). + eapply type_pack_inversion in Hty1 as (B' & C & [= <-] & Hty1 & ? & ?). + eapply typed_substitutivity. + { apply Hty1. } + rewrite fmap_empty in Hty2. + eapply typed_subst_type_closed'; eauto. + replace A with A.[ids] by by asimpl. + enough (A.[ids] = A.[ren (+1)]) as -> by done. + eapply type_wf_subst_dom; first done. intros; lia. + - (* unop *) + eapply unop_inversion in Hty as (A1 & Hop & Hty). + assert ((A1 = Int ∧ A = Int) ∨ (A1 = Bool ∧ A = Bool)) as [(-> & ->) | (-> & ->)]. + { inversion Hop; subst; eauto. } + + eapply canonical_values_int in Hty as [n ->]; last by eapply is_val_spec; eauto. + simpl in Heq. injection Heq as <-. + inversion Hop; subst; simpl in *; injection Heval as <-; constructor. + + eapply canonical_values_bool in Hty as [b ->]; last by eapply is_val_spec; eauto. + simpl in Heq. injection Heq as <-. + inversion Hop; subst; simpl in *; injection Heval as <-; constructor. + - (* binop *) + eapply binop_inversion in Hty as (A1 & A2 & Hop & Hty1 & Hty2). + assert (A1 = Int ∧ A2 = Int ∧ (A = Int ∨ A = Bool)) as (-> & -> & HC). + { inversion Hop; subst; eauto. } + eapply canonical_values_int in Hty1 as [n ->]; last by eapply is_val_spec; eauto. + eapply canonical_values_int in Hty2 as [m ->]; last by eapply is_val_spec; eauto. + simpl in Heq1, Heq2. injection Heq1 as <-. injection Heq2 as <-. + simpl in Heval. + inversion Hop; subst; simpl in *; injection Heval as <-; constructor. + - by eapply if_inversion in Hty as (H1 & H2 & H3). + - by eapply if_inversion in Hty as (H1 & H2 & H3). + - by eapply fst_inversion in Hty as (B & (? & ? & [= <- <-] & ? & ?)%pair_inversion). + - by eapply snd_inversion in Hty as (B & (? & ? & [= <- <-] & ? & ?)%pair_inversion). + - eapply case_inversion in Hty as (B & C & (? & ? & [= <- <-] & Hty & ?)%injl_inversion & ? & ?). + eauto. + - eapply case_inversion in Hty as (B & C & (? & ? & [= <- <-] & Hty & ?)%injr_inversion & ? & ?). + eauto. + - (* unroll *) + eapply unroll_inversion in Hty as (B & -> & Hty). + eapply roll_inversion in Hty as (C & Heq & Hty). injection Heq as ->. done. +Qed. + +Lemma typed_preservation e e' A: + TY 0; ∅ ⊢ e : A → + contextual_step e e' → + TY 0; ∅ ⊢ e' : A. +Proof. + intros Hty Hstep. destruct Hstep as [K e1 e2 -> -> Hstep]. + eapply fill_typing_decompose in Hty as [B [H1 H2]]. + eapply fill_typing_compose; last done. + by eapply typed_preservation_base_step. +Qed. + +Lemma typed_safety e1 e2 A: + TY 0; ∅ ⊢ e1 : A → + rtc contextual_step e1 e2 → + is_val e2 ∨ reducible e2. +Proof. + induction 2; eauto using typed_progress, typed_preservation. +Qed. + + +(** Derived typing rules *) +Lemma typed_unroll' n Γ e A B: + TY n; Γ ⊢ e : (μ: A) → + B = A.[(μ: A)%ty/] → + TY n; Γ ⊢ (unroll e) : B. +Proof. + intros ? ->. by eapply typed_unroll. +Qed. + +Lemma typed_tapp' n Γ A B C e : + TY n; Γ ⊢ e : (∀: A) → + type_wf n B → + C = A.[B/] → + TY n; Γ ⊢ e <> : C. +Proof. + intros; subst C; by eapply typed_tapp. +Qed. diff --git a/theories/type_systems/systemf_mu/untyped_encoding.v b/theories/type_systems/systemf_mu/untyped_encoding.v new file mode 100644 index 0000000..5c42452 --- /dev/null +++ b/theories/type_systems/systemf_mu/untyped_encoding.v @@ -0,0 +1,80 @@ +From stdpp Require Import gmap base relations tactics. +From iris Require Import prelude. +From semantics.ts.systemf_mu Require Import lang notation types pure tactics. +From Autosubst Require Import Autosubst. + +(** * Encoding of the untyped lambda calculus *) + +Definition D := (μ: #0 → #0)%ty. + +Definition lame (x : string) (e : expr) : val := RollV (λ: x, e). +Definition appe (e1 e2 : expr) : expr := (unroll e1) e2. + +Lemma lame_typed n Γ x e : + TY n; (<[x := D]> Γ) ⊢ e : D → + TY n; Γ ⊢ lame x e : D. +Proof. solve_typing. Qed. + +Lemma app_typed n Γ e1 e2 : + TY n; Γ ⊢ e1 : D → + TY n; Γ ⊢ e2 : D → + TY n; Γ ⊢ appe e1 e2 : D. +Proof. + solve_typing. +Qed. + +Lemma appe_step_l e1 e1' (v : val) : + contextual_step e1 e1' → + contextual_step (appe e1 v) (appe e1' v). +Proof. + intros Hstep. unfold appe. + by eapply (fill_contextual_step [UnrollCtx; AppLCtx _]). +Qed. + +Lemma appe_step_r e1 e2 e2' : + contextual_step e2 e2' → + contextual_step (appe e1 e2) (appe e1 e2'). +Proof. + intros Hstep. unfold appe. + by eapply (fill_contextual_step [AppRCtx _]). +Qed. + +Lemma lame_step_beta x e (v : val) : + rtc contextual_step (appe (lame x e) v) (lang.subst x v e). +Proof. + unfold appe, lame. + econstructor 2. + { eapply (fill_contextual_step [AppLCtx _]). + eapply base_contextual_step. by econstructor. + } + econstructor 2. + { eapply base_contextual_step. econstructor; eauto. } + reflexivity. +Qed. + +(* Divergence *) +Definition ω : expr := + roll (λ: "x", (unroll "x") "x"). + +Definition Ω := ((unroll ω) ω)%E. + +Lemma Ω_loops: + tc contextual_step Ω Ω. +Proof. + econstructor 2; [|econstructor 1]. + - rewrite /Ω /ω. eauto. + - rewrite /Ω. eauto. +Qed. + +Lemma ω_typed : + TY 0; ∅ ⊢ ω : D. +Proof. + solve_typing. +Qed. + +Lemma Ω_typed : TY 0; ∅ ⊢ Ω : D. +Proof. + rewrite /Ω. econstructor. + - eapply typed_unroll'; eauto using ω_typed. + - eapply ω_typed. +Qed.