work on soundness proofs & fix bug in translate_typing
This commit is contained in:
parent
dbfe0cf4de
commit
fd2c035902
2 changed files with 163 additions and 232 deletions
327
coq/soundness.v
327
coq/soundness.v
|
@ -22,9 +22,24 @@ Proof.
|
||||||
|
|
||||||
- apply T_Let with (σ:=σ0).
|
- apply T_Let with (σ:=σ0).
|
||||||
apply IHexpr_type1.
|
apply IHexpr_type1.
|
||||||
|
|
||||||
admit.
|
admit.
|
||||||
Admitted.
|
Admitted.
|
||||||
|
|
||||||
|
Lemma typing_subst : forall Γ x σ s e τ,
|
||||||
|
(ctx_assign x σ Γ) |- e \is τ ->
|
||||||
|
Γ |- s \is σ ->
|
||||||
|
Γ |- (expr_subst x s e) \is τ.
|
||||||
|
Proof.
|
||||||
|
Admitted.
|
||||||
|
|
||||||
|
Lemma typing_tsubst : forall Γ α σ e τ,
|
||||||
|
Γ |- e \is τ ->
|
||||||
|
Γ |- (expr_specialize α σ e) \is (type_subst α σ τ).
|
||||||
|
Proof.
|
||||||
|
Admitted.
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
Lemma map_type : forall Γ,
|
Lemma map_type : forall Γ,
|
||||||
Γ |- [{ %"map"% }] \is [<
|
Γ |- [{ %"map"% }] \is [<
|
||||||
∀"σ",∀"τ", (%"σ"% -> %"τ"%) -> [%"σ"%] -> [%"τ"%]
|
∀"σ",∀"τ", (%"σ"% -> %"τ"%) -> [%"σ"%] -> [%"τ"%]
|
||||||
|
@ -32,6 +47,17 @@ Lemma map_type : forall Γ,
|
||||||
Proof.
|
Proof.
|
||||||
Admitted.
|
Admitted.
|
||||||
|
|
||||||
|
Lemma specialized_map_type : forall Γ τ τ',
|
||||||
|
Γ |- [{ %"map"% # τ # τ' }] \is [<
|
||||||
|
(τ -> τ') -> [τ] -> [τ']
|
||||||
|
>].
|
||||||
|
Proof.
|
||||||
|
Admitted.
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
|
(* morphism has valid typing *)
|
||||||
Lemma morphism_path_solves_type : forall Γ τ τ' m,
|
Lemma morphism_path_solves_type : forall Γ τ τ' m,
|
||||||
(translate_morphism_path Γ τ τ' m) ->
|
(translate_morphism_path Γ τ τ' m) ->
|
||||||
Γ |- m \is (type_morph τ τ')
|
Γ |- m \is (type_morph τ τ')
|
||||||
|
@ -86,75 +112,100 @@ Proof.
|
||||||
apply T_App with (σ':=(type_spec (type_id "Seq") τ)) (σ:=(type_spec (type_id "Seq") τ)).
|
apply T_App with (σ':=(type_spec (type_id "Seq") τ)) (σ:=(type_spec (type_id "Seq") τ)).
|
||||||
apply T_App with (σ':=(type_fun τ τ')) (σ:=(type_fun τ τ')).
|
apply T_App with (σ':=(type_fun τ τ')) (σ:=(type_fun τ τ')).
|
||||||
|
|
||||||
set (k:=[< (%"σ"% -> %"τ"%) -> <$"Seq"$ %"σ"%> -> <$"Seq"$ %"τ"%> >]).
|
apply specialized_map_type.
|
||||||
set (k1:=[< (τ -> %"τ"%) -> <$"Seq"$ τ> -> <$"Seq"$ %"τ"%> >]).
|
|
||||||
set (k2:=[< (τ -> τ') -> <$"Seq"$ τ> -> <$"Seq"$ τ'> >]).
|
|
||||||
|
|
||||||
set (P:=(type_subst "τ" τ' k1) = k2).
|
|
||||||
|
|
||||||
(* apply T_TypeApp with (α:="τ"%string) (τ:=k2).*)
|
|
||||||
(* apply T_TypeApp with (α:="τ"%string) (τ:=(type_subst "τ" τ' k1)).*)
|
|
||||||
(*
|
|
||||||
apply map_type.
|
|
||||||
|
|
||||||
apply TSubst_UnivReplace.
|
|
||||||
admit.
|
|
||||||
admit.
|
|
||||||
|
|
||||||
apply TSubst_UnivReplace.
|
|
||||||
|
|
||||||
apply T_MorphFun.
|
|
||||||
apply typing_weakening.
|
apply typing_weakening.
|
||||||
|
apply T_MorphFun.
|
||||||
apply IHtranslate_morphism_path.
|
apply IHtranslate_morphism_path.
|
||||||
apply M_Sub, TSubRepr_Refl, TEq_Refl.
|
|
||||||
|
|
||||||
apply T_Var.
|
apply id_morphism_path.
|
||||||
apply C_take.
|
auto using T_Var, C_take.
|
||||||
apply M_Sub, TSubRepr_Refl, TEq_Refl.
|
apply id_morphism_path.
|
||||||
*)
|
Qed.
|
||||||
Admitted.
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
(* reduction step preserves well-typedness *)
|
(* reduction step preserves well-typedness *)
|
||||||
Lemma preservation : forall Γ e e' τ,
|
Lemma preservation : forall Γ e e' τ,
|
||||||
~(is_value e) ->
|
|
||||||
(Γ |- e \is τ) ->
|
(Γ |- e \is τ) ->
|
||||||
(e -->β e') ->
|
(e -->β e') ->
|
||||||
(Γ |- e' \is τ)
|
(Γ |- e' \is τ)
|
||||||
.
|
.
|
||||||
Proof.
|
Proof.
|
||||||
intros.
|
intros Γ e e' τ Typ Red.
|
||||||
induction e.
|
generalize dependent e'.
|
||||||
|
induction Typ; intros e' Red.
|
||||||
|
|
||||||
(* `e` is Variable *)
|
(* `e` is Variable *)
|
||||||
contradict H.
|
- inversion Red.
|
||||||
apply V_Abs, VAbs_Var.
|
|
||||||
|
(* `e` is Let *)
|
||||||
|
- inversion Red.
|
||||||
|
subst.
|
||||||
|
apply typing_subst with (σ:=σ).
|
||||||
|
auto.
|
||||||
|
auto.
|
||||||
|
|
||||||
(* `e` is Type-Abstraction *)
|
(* `e` is Type-Abstraction *)
|
||||||
contradict H.
|
- inversion Red.
|
||||||
apply V_Abs, VAbs_TypAbs.
|
subst.
|
||||||
|
apply T_TypeAbs.
|
||||||
|
auto.
|
||||||
|
|
||||||
(* `e` is Type-Application *)
|
(* `e` is Type-Application *)
|
||||||
|
- inversion Red.
|
||||||
admit.
|
admit.
|
||||||
|
(*
|
||||||
|
apply typing_tsubst.
|
||||||
|
admit.
|
||||||
|
*)
|
||||||
|
|
||||||
(* `e` is Abstraction *)
|
(* `e` is abstraction *)
|
||||||
contradict H.
|
- inversion Red.
|
||||||
apply V_Abs, VAbs_Abs.
|
|
||||||
|
|
||||||
(* `e` is morphism *)
|
(* `e` is morphism *)
|
||||||
contradict H.
|
- inversion Red.
|
||||||
apply V_Abs, VAbs_Morph.
|
|
||||||
|
|
||||||
(* `e` is Application *)
|
(* `e` is Application *)
|
||||||
admit.
|
- inversion Red.
|
||||||
|
|
||||||
(* `e` is Let-Binding *)
|
* apply T_App with (σ':=σ') (σ:=σ).
|
||||||
admit.
|
apply IHTyp1.
|
||||||
|
auto. auto. auto.
|
||||||
|
|
||||||
(* `e` is Ascension *)
|
* apply T_App with (σ':=σ') (σ:=σ).
|
||||||
|
auto using IHTyp2.
|
||||||
|
auto. auto.
|
||||||
|
|
||||||
|
* apply typing_subst with (σ:=σ').
|
||||||
|
subst.
|
||||||
admit.
|
admit.
|
||||||
|
auto.
|
||||||
|
|
||||||
|
* apply typing_subst with (σ:=σ').
|
||||||
|
admit.
|
||||||
|
auto.
|
||||||
|
|
||||||
|
(* `e` is Morphism *)
|
||||||
|
- auto using T_MorphFun.
|
||||||
|
|
||||||
|
(* `e` is Ascend *)
|
||||||
|
- admit.
|
||||||
|
(*
|
||||||
|
apply IHexpr_type.
|
||||||
|
inversion Red.
|
||||||
|
*)
|
||||||
|
|
||||||
(* `e` is Descension *)
|
(* `e` is Descension *)
|
||||||
|
- intros.
|
||||||
|
apply T_DescendImplicit with (τ:=τ).
|
||||||
|
auto. auto.
|
||||||
|
|
||||||
|
(* `e` is descension *)
|
||||||
|
- apply T_DescendImplicit with (τ:=τ).
|
||||||
|
apply IHTyp.
|
||||||
admit.
|
admit.
|
||||||
|
auto.
|
||||||
Admitted.
|
Admitted.
|
||||||
|
|
||||||
(* translation of expression preserves typing *)
|
(* translation of expression preserves typing *)
|
||||||
|
@ -164,170 +215,46 @@ Lemma translation_preservation : forall Γ e e' τ,
|
||||||
(Γ |- e' \is τ)
|
(Γ |- e' \is τ)
|
||||||
.
|
.
|
||||||
Proof.
|
Proof.
|
||||||
intros.
|
intros Γ e e' τ Typ Transl.
|
||||||
induction H0.
|
generalize dependent e'.
|
||||||
|
intros e' Transl.
|
||||||
|
induction Transl.
|
||||||
|
|
||||||
(* e is Variable *)
|
(* e is Variable *)
|
||||||
apply H.
|
- apply H.
|
||||||
|
|
||||||
(* e is Let-Binding *)
|
(* e is Let-Binding *)
|
||||||
apply T_Let with (τ:=τ) (σ:=σ).
|
- apply T_Let with (τ:=τ) (σ:=σ).
|
||||||
apply IHtranslate_typing1.
|
auto. auto.
|
||||||
apply H0.
|
|
||||||
apply IHtranslate_typing2.
|
|
||||||
apply H1.
|
|
||||||
|
|
||||||
(* e is Type-Abstraction *)
|
- auto using T_TypeAbs.
|
||||||
apply T_TypeAbs.
|
|
||||||
apply IHtranslate_typing.
|
|
||||||
apply H0.
|
|
||||||
|
|
||||||
(* e is Type-Application *)
|
(* e is Type-Application *)
|
||||||
|
- apply T_TypeApp.
|
||||||
admit.
|
admit.
|
||||||
|
|
||||||
(* e is Abstraction *)
|
- auto using T_Abs.
|
||||||
apply T_Abs.
|
- auto using T_MorphAbs.
|
||||||
apply IHtranslate_typing.
|
|
||||||
apply H0.
|
|
||||||
|
|
||||||
(* e is Morphism-Abstraction *)
|
|
||||||
apply T_MorphAbs.
|
|
||||||
apply IHtranslate_typing.
|
|
||||||
apply H0.
|
|
||||||
|
|
||||||
(* e is Application *)
|
(* e is Application *)
|
||||||
apply T_App with (σ':=σ) (σ:=σ) (τ:=τ).
|
- apply T_App with (σ':=σ) (σ:=σ) (τ:=τ).
|
||||||
apply IHtranslate_typing1.
|
auto.
|
||||||
apply H0.
|
apply T_App with (σ':=σ') (σ:=σ') (τ:=σ).
|
||||||
|
|
||||||
induction H3.
|
|
||||||
|
|
||||||
(* Repr-Subtype *)
|
|
||||||
apply T_App with (σ':=τ0) (σ:=τ0) (τ:=τ').
|
|
||||||
apply T_MorphFun.
|
apply T_MorphFun.
|
||||||
apply T_MorphAbs.
|
|
||||||
apply T_Descend with (τ:=τ0).
|
|
||||||
apply T_Var.
|
|
||||||
apply C_take.
|
|
||||||
apply H3.
|
|
||||||
|
|
||||||
apply T_DescendImplicit with (τ:=τ0).
|
|
||||||
apply IHtranslate_typing2.
|
|
||||||
apply H1.
|
|
||||||
apply TSubRepr_Refl, TEq_Refl.
|
|
||||||
apply M_Sub, TSubRepr_Refl, TEq_Refl.
|
|
||||||
|
|
||||||
(* Lifted Morphism *)
|
|
||||||
apply T_App with (σ':=(type_ladder σ τ0)) (σ:=(type_ladder σ τ0)) (τ:=(type_ladder σ τ')).
|
|
||||||
apply T_MorphFun.
|
|
||||||
apply T_MorphAbs.
|
|
||||||
apply T_Ascend with (τ:=τ').
|
|
||||||
apply T_App with (σ':=τ0) (σ:=τ0) (τ:=τ').
|
|
||||||
apply T_MorphFun.
|
|
||||||
apply typing_weakening.
|
|
||||||
apply morphism_path_solves_type.
|
apply morphism_path_solves_type.
|
||||||
apply H4.
|
auto. auto.
|
||||||
|
apply id_morphism_path.
|
||||||
apply T_Descend with (τ:=(type_ladder σ τ0)).
|
apply id_morphism_path.
|
||||||
apply T_Var.
|
|
||||||
apply C_take.
|
|
||||||
|
|
||||||
apply TSubRepr_Ladder, TSubRepr_Refl, TEq_Refl.
|
|
||||||
apply M_Sub, TSubRepr_Refl, TEq_Refl.
|
|
||||||
apply IHtranslate_typing2.
|
|
||||||
apply H1.
|
|
||||||
|
|
||||||
apply M_Sub, TSubRepr_Refl, TEq_Refl.
|
|
||||||
|
|
||||||
(* argument coecrion is single function variable *)
|
|
||||||
apply T_App with (σ':= τ0) (σ:=τ0).
|
|
||||||
apply T_MorphFun.
|
|
||||||
apply T_Var.
|
|
||||||
apply H3.
|
|
||||||
apply IHtranslate_typing2.
|
|
||||||
apply H1.
|
|
||||||
(* lemma: every context implies identity morphism *)
|
|
||||||
apply M_Sub, TSubRepr_Refl, TEq_Refl.
|
|
||||||
|
|
||||||
(* argument coecrion is chain of coercions *)
|
|
||||||
apply T_App with (σ':=τ0) (σ:=τ0).
|
|
||||||
apply T_MorphFun.
|
|
||||||
apply T_MorphAbs.
|
|
||||||
apply T_App with (σ':=τ') (σ:=τ').
|
|
||||||
apply T_MorphFun.
|
|
||||||
apply typing_weakening.
|
|
||||||
apply morphism_path_solves_type.
|
|
||||||
apply H3_0.
|
|
||||||
apply T_App with (σ':=τ0) (σ:=τ0).
|
|
||||||
apply T_MorphFun.
|
|
||||||
apply typing_weakening.
|
|
||||||
apply morphism_path_solves_type.
|
|
||||||
apply H3_.
|
|
||||||
apply T_Var.
|
|
||||||
apply C_take.
|
|
||||||
(* lemma: every context implies identity morphism *)
|
|
||||||
apply M_Sub, TSubRepr_Refl, TEq_Refl.
|
|
||||||
(* lemma: every context implies identity morphism *)
|
|
||||||
apply M_Sub, TSubRepr_Refl, TEq_Refl.
|
|
||||||
apply IHtranslate_typing2.
|
|
||||||
apply H1.
|
|
||||||
(* lemma: every context implies identity morphism *)
|
|
||||||
apply M_Sub, TSubRepr_Refl, TEq_Refl.
|
|
||||||
|
|
||||||
(* argument coercion is is map *)
|
|
||||||
(*
|
|
||||||
apply T_App with (σ':=(type_spec (type_id "Seq") τ0)) (σ:=(type_spec (type_id "Seq") τ0)).
|
|
||||||
apply T_MorphFun.
|
|
||||||
apply T_MorphAbs.
|
|
||||||
apply T_App with (σ':=(type_spec (type_id "Seq") τ0)) (σ:=(type_spec (type_id "Seq") τ0)).
|
|
||||||
apply T_App with (σ':=(type_fun τ0 τ')) (σ:=(type_fun τ0 τ')).
|
|
||||||
apply T_TypeApp with
|
|
||||||
(α:="T2"%string)
|
|
||||||
(e:=(expr_ty_app (expr_var "map") τ0))
|
|
||||||
(τ:=(type_fun
|
|
||||||
(type_fun τ0 τ')
|
|
||||||
(type_fun
|
|
||||||
(type_spec (type_id "Seq") τ0)
|
|
||||||
(type_spec (type_id "Seq") τ')))).
|
|
||||||
|
|
||||||
apply T_TypeApp with
|
|
||||||
(α:="T1"%string)
|
|
||||||
(e:=(expr_var "map"))
|
|
||||||
(τ:=(type_univ "T2"
|
|
||||||
(type_fun
|
|
||||||
(type_fun τ0 τ')
|
|
||||||
(type_fun
|
|
||||||
(type_spec (type_id "Seq") (type_var "T1"))
|
|
||||||
(type_spec (type_id "Seq") (type_var "T2")))))).
|
|
||||||
|
|
||||||
apply T_Var.
|
|
||||||
admit.
|
|
||||||
|
|
||||||
apply TSubst_VarReplace.
|
|
||||||
apply TSubst_UnivReplace.
|
|
||||||
*)
|
|
||||||
admit.
|
|
||||||
|
|
||||||
(* argument coercion *)
|
|
||||||
apply M_Sub, TSubRepr_Refl, TEq_Refl.
|
|
||||||
|
|
||||||
(* end case `e application` *)
|
|
||||||
|
|
||||||
(* e is Morphism *)
|
(* e is Morphism *)
|
||||||
apply T_MorphFun.
|
- auto using T_MorphFun.
|
||||||
apply IHtranslate_typing.
|
|
||||||
apply H0.
|
|
||||||
|
|
||||||
(* e is Ascension *)
|
(* e is Ascension *)
|
||||||
apply T_Ascend.
|
- auto using T_Ascend.
|
||||||
apply IHtranslate_typing.
|
|
||||||
apply H0.
|
|
||||||
|
|
||||||
(* e is Desecension *)
|
(* e is Desecension *)
|
||||||
apply T_Descend with (τ:=τ).
|
- apply T_Descend with (τ:=τ).
|
||||||
apply IHtranslate_typing.
|
auto. auto.
|
||||||
apply H0.
|
|
||||||
apply H1.
|
|
||||||
Admitted.
|
Admitted.
|
||||||
|
|
||||||
(* e is stuck when it is neither a value, nor can it be reduced *)
|
(* e is stuck when it is neither a value, nor can it be reduced *)
|
||||||
|
@ -346,6 +273,8 @@ Lemma progress :
|
||||||
Proof.
|
Proof.
|
||||||
Admitted.
|
Admitted.
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
(* every well-typed expression is translated,
|
(* every well-typed expression is translated,
|
||||||
* such that it be reduced to a value
|
* such that it be reduced to a value
|
||||||
*)
|
*)
|
||||||
|
@ -357,20 +286,21 @@ Theorem soundness :
|
||||||
.
|
.
|
||||||
Proof.
|
Proof.
|
||||||
intros.
|
intros.
|
||||||
|
induction H0.
|
||||||
|
|
||||||
(* `e` is Variable *)
|
(* `e` is Variable *)
|
||||||
induction H0.
|
- exists (expr_var x).
|
||||||
exists (expr_var x).
|
|
||||||
split. apply Multi_Refl.
|
split. apply Multi_Refl.
|
||||||
split. apply V_Abs,VAbs_Var.
|
split. apply V_Abs,VAbs_Var.
|
||||||
apply H.
|
apply H.
|
||||||
|
|
||||||
(* `e` is Let-Binding *)
|
(* `e` is Let-Binding *)
|
||||||
exists (expr_subst x e' t').
|
- exists (expr_subst x e' t').
|
||||||
split.
|
split.
|
||||||
apply Multi_Step with (y:=(expr_subst x e' t')).
|
apply Multi_Step with (y:=(expr_subst x e' t')).
|
||||||
apply E_Let with (x:=x) (a:=e') (e:=t').
|
apply E_Let with (x:=x) (a:=e') (e:=t').
|
||||||
apply Multi_Refl.
|
apply Multi_Refl.
|
||||||
|
admit.
|
||||||
(*
|
(*
|
||||||
split.
|
split.
|
||||||
unfold expr_subst.
|
unfold expr_subst.
|
||||||
|
@ -381,10 +311,9 @@ Proof.
|
||||||
unfold expr_subst.
|
unfold expr_subst.
|
||||||
apply E_Let.
|
apply E_Let.
|
||||||
*)
|
*)
|
||||||
admit.
|
|
||||||
|
|
||||||
(* `e` is Type-Abstraction *)
|
(* `e` is Type-Abstraction *)
|
||||||
exists (expr_ty_abs α e').
|
- exists (expr_ty_abs α e').
|
||||||
split.
|
split.
|
||||||
apply Multi_Refl.
|
apply Multi_Refl.
|
||||||
split.
|
split.
|
||||||
|
@ -395,10 +324,10 @@ Proof.
|
||||||
apply H1.
|
apply H1.
|
||||||
|
|
||||||
(* `e` is Type-Application *)
|
(* `e` is Type-Application *)
|
||||||
admit.
|
- admit.
|
||||||
|
|
||||||
(* `e`is Abstraction *)
|
(* `e`is Abstraction *)
|
||||||
exists (expr_abs x σ e').
|
- exists (expr_abs x σ e').
|
||||||
split. apply Multi_Refl.
|
split. apply Multi_Refl.
|
||||||
split. apply V_Abs, VAbs_Abs.
|
split. apply V_Abs, VAbs_Abs.
|
||||||
apply T_Abs.
|
apply T_Abs.
|
||||||
|
@ -407,7 +336,7 @@ Proof.
|
||||||
apply H2.
|
apply H2.
|
||||||
|
|
||||||
(* `e` is Morphism Abstraction *)
|
(* `e` is Morphism Abstraction *)
|
||||||
exists (expr_morph x σ e').
|
- exists (expr_morph x σ e').
|
||||||
split. apply Multi_Refl.
|
split. apply Multi_Refl.
|
||||||
split. apply V_Abs, VAbs_Morph.
|
split. apply V_Abs, VAbs_Morph.
|
||||||
apply T_MorphAbs.
|
apply T_MorphAbs.
|
||||||
|
@ -416,13 +345,15 @@ Proof.
|
||||||
apply H2.
|
apply H2.
|
||||||
|
|
||||||
(* `e` is Application *)
|
(* `e` is Application *)
|
||||||
admit.
|
- admit.
|
||||||
admit.
|
|
||||||
|
(* `e` is morphism *)
|
||||||
|
- admit.
|
||||||
|
|
||||||
(* `e` is Ascension *)
|
(* `e` is Ascension *)
|
||||||
admit.
|
- admit.
|
||||||
|
|
||||||
(* `e` is Descension *)
|
(* `e` is Descension *)
|
||||||
admit.
|
- admit.
|
||||||
Admitted.
|
Admitted.
|
||||||
|
|
||||||
|
|
|
@ -101,9 +101,9 @@ Inductive translate_typing : context -> expr_term -> type_term -> expr_term -> P
|
||||||
(Γ |- e \is (type_univ α τ)) ->
|
(Γ |- e \is (type_univ α τ)) ->
|
||||||
(translate_typing Γ e τ e') ->
|
(translate_typing Γ e τ e') ->
|
||||||
(translate_typing Γ
|
(translate_typing Γ
|
||||||
[{ e # τ }]
|
[{ e # σ }]
|
||||||
(type_subst α σ τ)
|
(type_subst α σ τ)
|
||||||
[{ e' # τ }])
|
[{ e' # σ }])
|
||||||
|
|
||||||
| Expand_Abs : forall Γ x σ e e' τ,
|
| Expand_Abs : forall Γ x σ e e' τ,
|
||||||
((ctx_assign x σ Γ) |- e \is τ) ->
|
((ctx_assign x σ Γ) |- e \is τ) ->
|
||||||
|
|
Loading…
Reference in a new issue