|
|
|
@ -143,12 +143,13 @@ Proof.
|
|
|
|
|
- intros y e'. rewrite lookup_insert_Some.
|
|
|
|
|
intros [[-> <-]|[Hne Hlook]].
|
|
|
|
|
+ by eapply expr_rel_closed.
|
|
|
|
|
+ eapply IHsem_context_rel; last done.
|
|
|
|
|
+ eapply IHsem_context_rel.
|
|
|
|
|
exact Hlook.
|
|
|
|
|
Qed.
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
(* This is essentially an inversion lemma for 𝒢 *)
|
|
|
|
|
Lemma sem_context_rel_exprs Γ θ x A :
|
|
|
|
|
Lemma sem_context_rel_exprs {Γ θ x A} :
|
|
|
|
|
sem_context_rel Γ θ →
|
|
|
|
|
Γ !! x = Some A →
|
|
|
|
|
∃ e, θ !! x = Some e ∧ ℰ A e.
|
|
|
|
@ -170,15 +171,249 @@ Proof.
|
|
|
|
|
- rewrite !dom_insert. congruence.
|
|
|
|
|
Qed.
|
|
|
|
|
|
|
|
|
|
Ltac specialize_all' θ H_ctx := repeat (
|
|
|
|
|
match goal with
|
|
|
|
|
| [ H: ∀ θ, ∀ (_: sem_context_rel ?G θ), ?e |- ?goal ] =>
|
|
|
|
|
specialize (H θ H_ctx)
|
|
|
|
|
| [ H: ?G ⊨ ?e : ?T |- ?goal ] =>
|
|
|
|
|
destruct H as [? H];
|
|
|
|
|
specialize (H θ H_ctx)
|
|
|
|
|
end
|
|
|
|
|
).
|
|
|
|
|
|
|
|
|
|
Ltac specialize_all_closed := repeat (
|
|
|
|
|
match goal with
|
|
|
|
|
| [ H: ?G ⊨ ?e : ?T |- ?goal ] =>
|
|
|
|
|
destruct H as [H _]
|
|
|
|
|
end
|
|
|
|
|
).
|
|
|
|
|
|
|
|
|
|
Ltac specialize_all θ H_ctx :=
|
|
|
|
|
match goal with
|
|
|
|
|
| [ |- ∀ θ, ∀ (_: sem_context_rel ?G θ), ?e ] => intros θ H_ctx; specialize_all' θ H_ctx
|
|
|
|
|
| [ |- ?G ⊨ ?e : ?T ] => econstructor; first (specialize_all_closed); last (intros θ H_ctx; specialize_all' θ H_ctx)
|
|
|
|
|
end.
|
|
|
|
|
|
|
|
|
|
Ltac break_expr_rel H val Hcl Hval :=
|
|
|
|
|
simp type_interp in H;
|
|
|
|
|
destruct H as (val & H & Hcl & Hval).
|
|
|
|
|
|
|
|
|
|
Ltac split3 := split; last split.
|
|
|
|
|
|
|
|
|
|
(* Search (?A !! ?B = Some ?C) (?B ∈ ?D). *)
|
|
|
|
|
(* Search (?x ∈ elements ?y). *)
|
|
|
|
|
|
|
|
|
|
Lemma compat_var Γ x A:
|
|
|
|
|
Γ !! x = Some A →
|
|
|
|
|
Γ ⊨ (Var x): A.
|
|
|
|
|
Proof.
|
|
|
|
|
intro H_some.
|
|
|
|
|
split.
|
|
|
|
|
- simpl.
|
|
|
|
|
apply bool_decide_pack.
|
|
|
|
|
rewrite elem_of_elements.
|
|
|
|
|
exact (elem_of_dom_2 Γ x A H_some).
|
|
|
|
|
- specialize_all θ H_ctx.
|
|
|
|
|
destruct (sem_context_rel_exprs H_ctx H_some) as (e & Hθ_some & He).
|
|
|
|
|
simpl.
|
|
|
|
|
rewrite Hθ_some.
|
|
|
|
|
exact He.
|
|
|
|
|
Qed.
|
|
|
|
|
|
|
|
|
|
Search (elem_of ?x (cons ?x ?y)).
|
|
|
|
|
|
|
|
|
|
(* Compatibility for [lam] unfortunately needs a very technical helper lemma. *)
|
|
|
|
|
Lemma lam_closed Γ θ (x : string) A e :
|
|
|
|
|
closed (elements (dom (<[x:=A]> Γ))) e →
|
|
|
|
|
𝒢 Γ θ →
|
|
|
|
|
closed [] (Lam x (subst_map (delete x θ) e)).
|
|
|
|
|
Proof.
|
|
|
|
|
intros Hcl Hctxt.
|
|
|
|
|
eapply subst_map_closed'_2.
|
|
|
|
|
- eapply closed_weaken; first done.
|
|
|
|
|
rewrite dom_delete dom_insert (sem_context_rel_dom Γ θ) //.
|
|
|
|
|
(* The [set_solver] tactic is great for solving goals involving set
|
|
|
|
|
inclusion and union. However, when set difference is involved, it can't
|
|
|
|
|
always solve the goal -- we need to help it by doing a case distinction on
|
|
|
|
|
whether the element we are considering is [x] or not. *)
|
|
|
|
|
intros y. destruct (decide (x = y)); set_solver.
|
|
|
|
|
- eapply subst_closed_weaken, sem_context_rel_closed; last done.
|
|
|
|
|
+ set_solver.
|
|
|
|
|
+ apply map_delete_subseteq.
|
|
|
|
|
Qed.
|
|
|
|
|
|
|
|
|
|
Lemma lam_closed' {Γ θ x A e}:
|
|
|
|
|
closed (elements (dom (<[x:=A]> Γ))) e →
|
|
|
|
|
𝒢 Γ θ →
|
|
|
|
|
closed [x] (subst_map (delete x θ) e).
|
|
|
|
|
Proof.
|
|
|
|
|
intros Hcl H_ctx.
|
|
|
|
|
eapply closed_subst_weaken.
|
|
|
|
|
eapply subst_closed_weaken; first reflexivity.
|
|
|
|
|
by apply map_delete_subseteq.
|
|
|
|
|
exact (sem_context_rel_closed _ _ H_ctx).
|
|
|
|
|
2: exact Hcl.
|
|
|
|
|
|
|
|
|
|
rewrite dom_insert.
|
|
|
|
|
intros y Hy_in Hy_notin.
|
|
|
|
|
|
|
|
|
|
rewrite (sem_context_rel_dom _ _ H_ctx) in Hy_in.
|
|
|
|
|
rewrite dom_delete in Hy_notin.
|
|
|
|
|
rewrite not_elem_of_difference in Hy_notin.
|
|
|
|
|
rewrite elem_of_elements in Hy_in.
|
|
|
|
|
set_solver.
|
|
|
|
|
Qed.
|
|
|
|
|
the internship
|
|
|
|
|
Lemma compat_lam Γ x e A B:
|
|
|
|
|
(<[x:=A]> Γ ⊢ e : B)%ty →
|
|
|
|
|
<[x:=A]> Γ ⊨ e : B →
|
|
|
|
|
Γ ⊨ (λ: (BNamed x), e) : (A → B).
|
|
|
|
|
Proof.
|
|
|
|
|
intros He_ty IHe.
|
|
|
|
|
|
|
|
|
|
specialize_all θ H_ctx.
|
|
|
|
|
- simpl.
|
|
|
|
|
rename IHe into IHcl.
|
|
|
|
|
eapply closed_weaken.
|
|
|
|
|
exact IHcl.
|
|
|
|
|
|
|
|
|
|
rewrite dom_insert.
|
|
|
|
|
|
|
|
|
|
induction (decide (x ∈ (dom Γ))) as [Hin | Hnotin]; set_solver.
|
|
|
|
|
- destruct IHe as [Hcl IHe].
|
|
|
|
|
|
|
|
|
|
simp type_interp.
|
|
|
|
|
eexists.
|
|
|
|
|
simpl; split3.
|
|
|
|
|
{
|
|
|
|
|
constructor.
|
|
|
|
|
}
|
|
|
|
|
{
|
|
|
|
|
by eapply lam_closed.
|
|
|
|
|
}
|
|
|
|
|
|
|
|
|
|
simp type_interp.
|
|
|
|
|
exists x.
|
|
|
|
|
eexists.
|
|
|
|
|
simpl; split3.
|
|
|
|
|
{
|
|
|
|
|
reflexivity.
|
|
|
|
|
}
|
|
|
|
|
{
|
|
|
|
|
by eapply lam_closed'.
|
|
|
|
|
}
|
|
|
|
|
intros e_arg He_arg.
|
|
|
|
|
|
|
|
|
|
rewrite subst_subst_map.
|
|
|
|
|
{
|
|
|
|
|
eapply IHe.
|
|
|
|
|
econstructor.
|
|
|
|
|
assumption.
|
|
|
|
|
assumption.
|
|
|
|
|
}
|
|
|
|
|
{
|
|
|
|
|
exact (sem_context_rel_closed _ _ H_ctx).
|
|
|
|
|
}
|
|
|
|
|
Qed.
|
|
|
|
|
|
|
|
|
|
Lemma compat_int Γ z:
|
|
|
|
|
Γ ⊨ LitInt z : Int.
|
|
|
|
|
Proof.
|
|
|
|
|
econstructor.
|
|
|
|
|
1: eauto.
|
|
|
|
|
simpl; intros; simp type_interp.
|
|
|
|
|
exists z.
|
|
|
|
|
split; last split; eauto.
|
|
|
|
|
simp type_interp.
|
|
|
|
|
exists z; reflexivity.
|
|
|
|
|
Qed.
|
|
|
|
|
|
|
|
|
|
Lemma compat_plus Γ e1 e2:
|
|
|
|
|
Γ ⊨ e1 : Int →
|
|
|
|
|
Γ ⊨ e2 : Int →
|
|
|
|
|
Γ ⊨ (e1 + e2) : Int.
|
|
|
|
|
Proof.
|
|
|
|
|
intros H1 H2.
|
|
|
|
|
destruct H1 as [Hcl1 He1].
|
|
|
|
|
destruct H2 as [Hcl2 He2].
|
|
|
|
|
econstructor.
|
|
|
|
|
1: naive_solver.
|
|
|
|
|
specialize_all θ H_ctx.
|
|
|
|
|
break_expr_rel He1 v1 Hcl1' Hv1.
|
|
|
|
|
break_expr_rel He2 v2 Hcl2' Hv2.
|
|
|
|
|
simp type_interp in Hv1; destruct Hv1 as [z1 ->].
|
|
|
|
|
simp type_interp in Hv2; destruct Hv2 as [z2 ->].
|
|
|
|
|
|
|
|
|
|
simpl.
|
|
|
|
|
simp type_interp.
|
|
|
|
|
exists (z1 + z2)%Z.
|
|
|
|
|
split3.
|
|
|
|
|
- econstructor; assumption.
|
|
|
|
|
- naive_solver.
|
|
|
|
|
- simp type_interp.
|
|
|
|
|
eauto.
|
|
|
|
|
Qed.
|
|
|
|
|
|
|
|
|
|
Lemma compat_app Γ e1 e2 A B:
|
|
|
|
|
Γ ⊨ e1 : (A → B) →
|
|
|
|
|
Γ ⊨ e2 : A →
|
|
|
|
|
Γ ⊨ App e1 e2 : B.
|
|
|
|
|
Proof.
|
|
|
|
|
intros Hsem_fn Hsem_e2.
|
|
|
|
|
|
|
|
|
|
specialize_all θ H_ctx.
|
|
|
|
|
- simpl; rewrite andb_True; split; assumption.
|
|
|
|
|
- simpl.
|
|
|
|
|
simp type_interp; simpl; eauto.
|
|
|
|
|
break_expr_rel Hsem_fn v_fn Hcl_fn Hv_fn.
|
|
|
|
|
break_expr_rel Hsem_e2 v_e2 Hcl_e2 Hv_e2.
|
|
|
|
|
simp type_interp in Hv_fn.
|
|
|
|
|
simpl in Hv_fn; destruct Hv_fn as (x & e_body & -> & Hcl_e & IHe_subst).
|
|
|
|
|
|
|
|
|
|
assert (ℰ A (subst_map θ e2)) as Hsem_subst.
|
|
|
|
|
{
|
|
|
|
|
simp type_interp.
|
|
|
|
|
eauto.
|
|
|
|
|
}
|
|
|
|
|
specialize (IHe_subst (subst_map θ e2) Hsem_subst).
|
|
|
|
|
break_expr_rel IHe_subst v_target Hcl_target Hv_target.
|
|
|
|
|
|
|
|
|
|
exists v_target; split3.
|
|
|
|
|
{
|
|
|
|
|
econstructor.
|
|
|
|
|
exact v_target.
|
|
|
|
|
exact Hsem_fn.
|
|
|
|
|
exact IHe_subst.
|
|
|
|
|
}
|
|
|
|
|
apply andb_True; split; eauto.
|
|
|
|
|
assumption.
|
|
|
|
|
Qed.
|
|
|
|
|
|
|
|
|
|
Lemma sem_soundness {Γ e A}:
|
|
|
|
|
(Γ ⊢ e: A)%ty →
|
|
|
|
|
Γ ⊨ e: A.
|
|
|
|
|
Proof.
|
|
|
|
|
induction 1.
|
|
|
|
|
- by apply compat_var.
|
|
|
|
|
- by apply compat_lam.
|
|
|
|
|
- by apply compat_int.
|
|
|
|
|
- by eapply compat_app.
|
|
|
|
|
- by apply compat_plus.
|
|
|
|
|
Qed.
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
Lemma termination e A :
|
|
|
|
|
(∅ ⊢ e : A)%ty →
|
|
|
|
|
∃ v, big_step e v.
|
|
|
|
|
Proof.
|
|
|
|
|
(* You may want to add suitable intermediate lemmas, like we did for the cbv
|
|
|
|
|
logical relation as seen in the lecture. *)
|
|
|
|
|
(* TODO: exercise *)
|
|
|
|
|
Admitted.
|
|
|
|
|
|
|
|
|
|
intro H_step.
|
|
|
|
|
remember (sem_soundness H_step) as H_sem.
|
|
|
|
|
destruct H_sem as [H_closed H_e].
|
|
|
|
|
clear HeqH_sem.
|
|
|
|
|
specialize (H_e ∅ sem_context_rel_empty).
|
|
|
|
|
simp type_interp in H_e.
|
|
|
|
|
destruct H_e as (target & Hbs_subst & _ & _).
|
|
|
|
|
exists target.
|
|
|
|
|
rewrite subst_map_empty in Hbs_subst.
|
|
|
|
|
assumption.
|
|
|
|
|
Qed.
|
|
|
|
|