diff --git a/coq/soundness.v b/coq/soundness.v index e550a7d..5558dde 100644 --- a/coq/soundness.v +++ b/coq/soundness.v @@ -16,16 +16,21 @@ Proof. intros. induction H. - apply T_Var. - apply C_shuffle. - apply H. + - apply T_Var. + apply C_shuffle. + apply H. - apply T_Let with (σ:=σ0). - apply IHexpr_type1. - admit. - + - apply T_Let with (σ:=σ0). + apply IHexpr_type1. + + admit. +Admitted. +Lemma map_type : forall Γ, + Γ |- [{ %"map"% }] \is [< + ∀"σ",∀"τ", (%"σ"% -> %"τ"%) -> [%"σ"%] -> [%"τ"%] + >]. +Proof. Admitted. - Lemma morphism_path_solves_type : forall Γ τ τ' m, (translate_morphism_path Γ τ τ' m) -> @@ -37,7 +42,7 @@ Proof. (* Sub *) apply T_MorphAbs. - apply T_DescendImplicit with (τ:=τ). + apply T_Descend with (τ:=τ). apply T_Var. apply C_take. apply H. @@ -53,7 +58,7 @@ Proof. apply T_Var. apply C_take. apply TSubRepr_Ladder, TSubRepr_Refl, TEq_Refl. - apply M_Sub, TSubRepr_Refl, TEq_Refl. + apply id_morphism_path. (* Single *) apply T_Var. @@ -65,7 +70,7 @@ Proof. apply T_MorphFun. apply typing_weakening. apply IHtranslate_morphism_path2. - + apply T_App with (σ':=τ) (σ:=τ) (τ:=τ'). apply T_MorphFun. apply typing_weakening. @@ -73,15 +78,30 @@ Proof. apply T_Var. apply C_take. - apply M_Sub, TSubRepr_Refl, TEq_Refl. - apply M_Sub, TSubRepr_Refl, TEq_Refl. + apply id_morphism_path. + apply id_morphism_path. (* Map Sequence *) apply T_MorphAbs. apply T_App with (σ':=(type_spec (type_id "Seq") τ)) (σ:=(type_spec (type_id "Seq") τ)). apply T_App with (σ':=(type_fun τ τ')) (σ:=(type_fun τ τ')). + set (k:=[< (%"σ"% -> %"τ"%) -> <$"Seq"$ %"σ"%> -> <$"Seq"$ %"τ"%> >]). + 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. @@ -91,6 +111,7 @@ Proof. apply T_Var. apply C_take. apply M_Sub, TSubRepr_Refl, TEq_Refl. + *) Admitted. (* reduction step preserves well-typedness *) @@ -185,7 +206,7 @@ Proof. apply T_App with (σ':=τ0) (σ:=τ0) (τ:=τ'). apply T_MorphFun. apply T_MorphAbs. - apply T_DescendImplicit with (τ:=τ0). + apply T_Descend with (τ:=τ0). apply T_Var. apply C_take. apply H3. @@ -303,7 +324,7 @@ Proof. apply H0. (* e is Desecension *) - apply T_DescendImplicit with (τ:=τ). + apply T_Descend with (τ:=τ). apply IHtranslate_typing. apply H0. apply H1. diff --git a/coq/subst.v b/coq/subst.v index 7cb13fc..5b8375f 100644 --- a/coq/subst.v +++ b/coq/subst.v @@ -1,60 +1,34 @@ - From Coq Require Import Strings.String. +From Coq Require Import Strings.String. +From Coq Require Import Lists.List. +Import ListNotations. Require Import terms. -(* Type Variable "x" is a free variable in type *) -Inductive type_var_free (x:string) : type_term -> Prop := - | TFree_Var : - (type_var_free x (type_var x)) - - | TFree_Ladder : forall τ1 τ2, - (type_var_free x τ1) -> - (type_var_free x τ2) -> - (type_var_free x (type_ladder τ1 τ2)) - - | TFree_Fun : forall τ1 τ2, - (type_var_free x τ1) -> - (type_var_free x τ2) -> - (type_var_free x (type_fun τ1 τ2)) - - | TFree_Morph : forall τ1 τ2, - (type_var_free x τ1) -> - (type_var_free x τ2) -> - (type_var_free x (type_morph τ1 τ2)) - - | TFree_Spec : forall τ1 τ2, - (type_var_free x τ1) -> - (type_var_free x τ2) -> - (type_var_free x (type_spec τ1 τ2)) - - | TFree_Univ : forall y τ, - ~(y = x) -> - (type_var_free x τ) -> - (type_var_free x (type_univ y τ)) -. +Fixpoint type_fv (τ : type_term) {struct τ} : (list string) := + match τ with + | type_id s => [] + | type_var α => [α] + | type_univ α τ => (remove string_dec α (type_fv τ)) + | type_spec σ τ => (type_fv σ) ++ (type_fv τ) + | type_fun σ τ => (type_fv σ) ++ (type_fv τ) + | type_morph σ τ => (type_fv σ) ++ (type_fv τ) + | type_ladder σ τ => (type_fv σ) ++ (type_fv τ) + end. Open Scope ladder_type_scope. -Example ex_type_free_var1 : - (type_var_free "T" (type_univ "U" (type_var "T"))) +Example ex_type_fv1 : + (In "T"%string (type_fv [< ∀"U",%"T"% >])) . -Proof. - apply TFree_Univ. - easy. - apply TFree_Var. -Qed. +Proof. simpl. left. auto. Qed. Open Scope ladder_type_scope. -Example ex_type_free_var2 : - ~(type_var_free "T" (type_univ "T" (type_var "T"))) +Example ex_type_fv2 : + ~(In "T"%string (type_fv [< ∀"T",%"T"% >])) . -Proof. - intro H. - inversion H. - contradiction. -Qed. +Proof. simpl. auto. Qed. (* scoped variable substitution in type terms $\label{coq:subst-type}$ *) -Fixpoint type_subst (v:string) (n:type_term) (t0:type_term) := +Fixpoint type_subst (v:string) (n:type_term) (t0:type_term) : type_term := match t0 with | type_var name => if (eqb v name) then n else t0 | type_fun t1 t2 => (type_fun (type_subst v n t1) (type_subst v n t2)) @@ -64,12 +38,14 @@ Fixpoint type_subst (v:string) (n:type_term) (t0:type_term) := | t => t end. +(* + Inductive type_subst1 (x:string) (σ:type_term) : type_term -> type_term -> Prop := | TSubst_VarReplace : (type_subst1 x σ (type_var x) σ) | TSubst_VarKeep : forall y, - ~(x = y) -> + (x <> y) -> (type_subst1 x σ (type_var y) (type_var y)) | TSubst_UnivReplace : forall y τ τ', @@ -101,6 +77,56 @@ Inductive type_subst1 (x:string) (σ:type_term) : type_term -> type_term -> Prop (type_subst1 x σ τ2 τ2') -> (type_subst1 x σ (type_ladder τ1 τ2) (type_ladder τ1' τ2')) . +*) + +Lemma type_subst_symm : + forall x y τ τ', + ((type_subst x (type_var y) τ) = τ') -> + ((type_subst y (type_var x) τ') = τ) +. +Proof. + intros. + induction H. + + unfold type_subst. + induction τ. + reflexivity. + + +Admitted. + +Lemma type_subst_fresh : + forall α τ u, + ~ (In α (type_fv τ)) + -> (type_subst α u τ) = τ +. +Proof. + intros. + unfold type_subst. + induction τ. + reflexivity. + + + unfold eqb. + + admit. +(* + apply TSubst_Id. + apply TSubst_VarKeep. + contradict H. + rewrite H. + apply TFree_Var. + + apply TSubst_Fun. + apply IHτ1. + + contradict H. + apply TFree_Fun. + apply H. + apply + *) +Admitted. + (* scoped variable substitution, replaces free occurences of v with n in expression e *) Fixpoint expr_subst (v:string) (n:expr_term) (e0:expr_term) := diff --git a/coq/typing.v b/coq/typing.v index 1777111..8f1a845 100644 --- a/coq/typing.v +++ b/coq/typing.v @@ -11,6 +11,7 @@ Require Import morph. (** Typing Derivation *) +Open Scope ladder_type_scope. Open Scope ladder_expr_scope. Reserved Notation "Gamma '|-' x '\is' X" (at level 101, x at next level, X at level 0). @@ -267,4 +268,116 @@ Proof. apply M_Sub. apply TSubRepr_Refl. apply TEq_Refl. Qed. -End Typing. + +Ltac var_typing := auto using T_Var, C_shuffle, C_take. +Ltac repr_subtype := auto using TSubRepr_Ladder, TSubRepr_Refl, TEq_Refl, TEq_LadderAssocLR. + +Example expand1 : + (translate_typing + (ctx_assign "60" [< $"ℝ"$ >] + (ctx_assign "360" [< $"ℝ"$ >] + (ctx_assign "/" [< $"ℝ"$ -> $"ℝ"$ -> $"ℝ"$ >] + ctx_empty))) + [{ + let "deg2turns" := + (λ"x",$"Angle"$~$"Degrees"$~$"ℝ"$ + ↦morph ((%"/"% (%"x"% des $"ℝ"$) %"360"%) as $"Angle"$~$"Turns"$)) in + let "sin" := (λ"α",$"Angle"$~$"Turns"$~$"ℝ"$ ↦ (%"α"% des $"ℝ"$)) in + ( %"sin"% (%"60"% as $"Angle"$~$"Degrees"$) ) + }] + [< + $"ℝ"$ + >] + [{ + let "deg2turns" := + (λ"x",$"Angle"$~$"Degrees"$~$"ℝ"$ + ↦morph ((%"/"% (%"x"% des $"ℝ"$) %"360"%) as $"Angle"$~$"Turns"$)) in + let "sin" := (λ"α",$"Angle"$~$"Turns"$~$"ℝ"$ ↦ (%"α"% des $"ℝ"$)) in + (%"sin"% (%"deg2turns"% (%"60"% as $"Angle"$~$"Degrees"$))) + }]) +. +Proof. + apply Expand_Let with (σ:=[< ($"Angle"$~$"Degrees"$)~$"ℝ"$ ->morph ($"Angle"$~$"Turns"$)~$"ℝ"$ >]). + + - apply T_DescendImplicit with (τ:=[< $"Angle"$~$"Degrees"$~$"ℝ"$ ->morph $"Angle"$~$"Turns"$ ~ $"ℝ"$ >]). + 2: { + apply TSubRepr_Refl. + apply TEq_SubMorph. + apply TEq_LadderAssocRL. + apply TEq_LadderAssocRL. + } + apply T_MorphAbs. + apply T_DescendImplicit with (τ:=[< ($"Angle"$~$"Turns"$) ~ $"ℝ"$ >]). + 2: { + apply TSubRepr_Refl. + apply TEq_LadderAssocLR. + } + apply T_Ascend with (τ:=[<$"ℝ"$>]) (τ':=[<$"Angle"$~$"Turns"$>]). + apply T_App with (σ':=[<$"ℝ"$>]) (σ:=[<$"ℝ"$>]). + apply T_App with (σ':=[<$"ℝ"$>]) (σ:=[<$"ℝ"$>]). + var_typing. + var_typing. + apply T_Descend with (τ:=[<$"Angle"$~$"Degrees"$~$"ℝ"$>]). + repr_subtype. + var_typing. + repr_subtype. + apply id_morphism_path. + var_typing. + apply id_morphism_path. + + - apply T_Let with (σ:=[< $"Angle"$~$"Turns"$~$"ℝ"$ -> $"ℝ"$ >]). + apply T_Abs. + * apply T_Descend with (τ:=[<$"Angle"$~$"Turns"$~$"ℝ"$>]). + 2: repr_subtype. + var_typing. + + * apply T_App with (σ':=[<($"Angle"$~$"Degrees"$)~$"ℝ"$>]) (σ:=[<($"Angle"$~$"Turns"$)~$"ℝ"$>]) (τ:=[<$"ℝ"$>]). + apply T_DescendImplicit with (τ:=[< $"Angle"$~$"Turns"$~$"ℝ"$ -> $"ℝ"$ >]). + 2: { + apply TSubRepr_Refl. + apply TEq_SubFun. + apply TEq_LadderAssocRL. + apply TEq_Refl. + } + var_typing. + + apply T_Ascend with (τ':=[<$"Angle"$~$"Degrees"$>]) (τ:=[<$"ℝ"$>]). + var_typing. + + apply M_Single with (h:="deg2turns"%string). + var_typing. + + - admit. + (* + - apply Expand_MorphAbs. + * apply T_DescendImplicit with (τ:=[< ($"Angle"$~$"Turns"$) ~ $"ℝ"$ >]). + 2: repr_subtype. + apply T_Ascend with (τ':=[<$"Angle"$~$"Turns"$>]) (τ:=[<$"ℝ"$>]). + apply T_App with (σ:=[<$"ℝ"$>]) (σ':=[<$"ℝ"$>]). + apply T_App with (σ:=[<$"ℝ"$>]) (σ':=[<$"ℝ"$>]). + auto using T_Var, C_shuffle, C_take. + apply T_Descend with (τ:=[<$"Angle"$~$"Degrees"$~$"ℝ"$>]). + 2: repr_subtype. + var_typing. + apply id_morphism_path. + var_typing. + apply id_morphism_path. + + * apply T_Abs. + apply T_DescendImplicit with (τ:=[< ($"Angle"$~$"Turns"$) ~ $"ℝ"$ >]). + 2: repr_subtype. + apply T_Ascend with (τ':=[<$"Angle"$~$"Turns"$>]) (τ:=[<$"ℝ"$>]). + apply T_App with (σ':=[<$"ℝ"$>]) (σ:=[<$"ℝ"$>]). + apply T_App with (σ':=[<$"ℝ"$>]) (σ:=[<$"ℝ"$>]). + var_typing. + apply T_Descend with [<$"Angle"$~$"Degrees"$~$"ℝ"$>]. + 2: repr_subtype. + var_typing. + apply id_morphism_path. + var_typing. + apply id_morphism_path. + + * apply Expand_Ascend. +*) + - admit. +Admitted.