ladder-calculus/coq/typing.v

384 lines
11 KiB
Coq
Raw Permalink Normal View History

(* This module defines the typing relation
* where each expression is assigned a type.
*)
From Coq Require Import Strings.String.
Require Import terms.
Require Import subst.
2024-08-21 20:04:20 +02:00
Require Import equiv.
Require Import subtype.
Require Import context.
Require Import morph.
2024-08-21 20:04:20 +02:00
(** Typing Derivation *)
Open Scope ladder_type_scope.
Open Scope ladder_expr_scope.
2024-07-27 13:30:12 +02:00
Reserved Notation "Gamma '|-' x '\is' X" (at level 101, x at next level, X at level 0).
2024-07-27 13:30:12 +02:00
Inductive expr_type : context -> expr_term -> type_term -> Prop :=
| T_Var : forall Γ x τ,
context_contains Γ x τ ->
Γ |- [{ %x% }] \is τ
2024-07-27 13:30:12 +02:00
| T_Let : forall Γ s σ t τ x,
Γ |- s \is σ ->
(ctx_assign x σ Γ) |- t \is τ ->
Γ |- [{ let x := s in t }] \is τ
2024-07-27 13:30:12 +02:00
| T_TypeAbs : forall Γ e τ α,
2024-07-27 13:30:12 +02:00
Γ |- e \is τ ->
Γ |- [{ Λ α e }] \is [< α,τ >]
2024-07-27 13:30:12 +02:00
| T_TypeApp : forall Γ α e σ τ,
Γ |- e \is [< α, τ >] ->
Γ |- [{ e # σ }] \is (type_subst α σ τ)
2024-07-27 13:30:12 +02:00
| T_Abs : forall Γ x σ t τ,
(ctx_assign x σ Γ) |- t \is τ ->
Γ |- [{ λ x , σ t }] \is [< σ -> τ >]
| T_MorphAbs : forall Γ x σ t τ,
(ctx_assign x σ Γ) |- t \is τ ->
Γ |- [{ λ x , σ morph t }] \is [< σ ->morph τ >]
2024-09-04 12:39:15 +02:00
| T_App : forall Γ f a σ' σ τ,
(Γ |- f \is [< σ -> τ >]) ->
(Γ |- a \is σ') ->
(Γ |- σ' ~> σ) ->
(Γ |- [{ (f a) }] \is τ)
2024-09-04 12:39:15 +02:00
| T_MorphFun : forall Γ f σ τ,
Γ |- f \is (type_morph σ τ) ->
Γ |- f \is (type_fun σ τ)
| T_Ascend : forall Γ e τ τ',
(Γ |- e \is τ) ->
(Γ |- [{ e as τ' }] \is [< τ' ~ τ >])
2024-09-04 12:39:15 +02:00
| T_DescendImplicit : forall Γ x τ τ',
2024-08-21 20:04:20 +02:00
Γ |- x \is τ ->
(τ :<= τ') ->
Γ |- x \is τ'
| T_Descend : forall Γ x τ τ',
Γ |- x \is τ ->
(τ :<= τ') ->
Γ |- [{ x des τ' }] \is τ'
2024-07-27 13:30:12 +02:00
where "Γ '|-' x '\is' τ" := (expr_type Γ x τ).
Definition is_well_typed (e:expr_term) : Prop :=
forall Γ,
exists τ,
2024-09-04 12:39:15 +02:00
Γ |- e \is τ
.
2024-09-08 15:34:15 +02:00
Inductive translate_typing : context -> expr_term -> type_term -> expr_term -> Prop :=
| Expand_Var : forall Γ x τ,
(Γ |- (expr_var x) \is τ) ->
(translate_typing Γ [{ %x% }] τ [{ %x% }])
2024-09-08 15:34:15 +02:00
| Expand_Let : forall Γ x e e' t t' σ τ,
(Γ |- e \is σ) ->
((ctx_assign x σ Γ) |- t \is τ) ->
(translate_typing Γ e σ e') ->
(translate_typing (ctx_assign x σ Γ) t τ t') ->
(translate_typing Γ
[{ let x := e in t }]
τ
[{ let x := e' in t' }])
2024-09-08 15:34:15 +02:00
| Expand_TypeAbs : forall Γ α e e' τ,
(Γ |- e \is τ) ->
(translate_typing Γ e τ e') ->
(translate_typing Γ
[{ Λ α e }]
[< α,τ >]
[{ Λ α e' }])
2024-09-08 15:34:15 +02:00
| Expand_TypeApp : forall Γ e e' σ α τ,
(Γ |- e \is (type_univ α τ)) ->
(translate_typing Γ e τ e') ->
(translate_typing Γ
[{ e # σ }]
(type_subst α σ τ)
[{ e' # σ }])
2024-09-08 15:34:15 +02:00
| Expand_Abs : forall Γ x σ e e' τ,
((ctx_assign x σ Γ) |- e \is τ) ->
(Γ |- (expr_abs x σ e) \is (type_fun σ τ)) ->
(translate_typing (ctx_assign x σ Γ) e τ e') ->
(translate_typing Γ
[{ λ x , σ e }]
[< σ -> τ >]
[{ λ x , σ e' }])
2024-09-08 15:34:15 +02:00
| Expand_MorphAbs : forall Γ x σ e e' τ,
((ctx_assign x σ Γ) |- e \is τ) ->
(Γ |- (expr_abs x σ e) \is (type_fun σ τ)) ->
(translate_typing (ctx_assign x σ Γ) e τ e') ->
(translate_typing Γ
[{ λ x , σ morph e }]
[< σ ->morph τ >]
[{ λ x , σ morph e' }])
2024-09-08 15:34:15 +02:00
| Expand_App : forall Γ f f' a a' m σ τ σ',
(Γ |- f \is (type_fun σ τ)) ->
(Γ |- a \is σ') ->
(Γ |- σ' ~> σ) ->
(translate_typing Γ f [< σ -> τ >] f') ->
2024-09-08 15:34:15 +02:00
(translate_typing Γ a σ' a') ->
(translate_morphism_path Γ σ' σ m) ->
(translate_typing Γ [{ f a }] τ [{ f' (m a') }])
2024-09-08 15:34:15 +02:00
| Expand_MorphFun : forall Γ f f' σ τ,
(Γ |- f \is (type_morph σ τ)) ->
(Γ |- f \is (type_fun σ τ)) ->
(translate_typing Γ f [< σ ->morph τ >] f') ->
(translate_typing Γ f [< σ -> τ >] f')
2024-09-08 15:34:15 +02:00
| Expand_Ascend : forall Γ e e' τ τ',
(Γ |- e \is τ) ->
(Γ |- [{ e as τ' }] \is [< τ' ~ τ >]) ->
2024-09-08 15:34:15 +02:00
(translate_typing Γ e τ e') ->
(translate_typing Γ [{ e as τ' }] [< τ' ~ τ >] [{ e' as τ' }])
2024-09-08 15:34:15 +02:00
| Expand_Descend : forall Γ e e' τ τ',
(Γ |- e \is τ) ->
(τ :<= τ') ->
(Γ |- [{ e des τ' }] \is τ') ->
2024-09-08 15:34:15 +02:00
(translate_typing Γ e τ e') ->
(translate_typing Γ e τ' [{ e' des τ' }])
2024-09-08 15:34:15 +02:00
.
2024-08-21 20:04:20 +02:00
(* Examples *)
Example typing1 :
ctx_empty |- [{ Λ"T" λ "x",%"T"% %"x"% }] \is [< "T", %"T"% -> %"T"% >].
2024-07-27 13:30:12 +02:00
Proof.
intros.
apply T_TypeAbs.
apply T_Abs.
apply T_Var.
apply C_take.
2024-08-21 20:04:20 +02:00
Qed.
2024-07-27 13:30:12 +02:00
Example typing2 :
ctx_empty |- [{ Λ"T" λ "x",%"T"% %"x"% }] \is [< "U", %"U"% -> %"U"% >].
Proof.
2024-08-21 20:04:20 +02:00
intros.
apply T_DescendImplicit with (τ:=[< "T",(%"T"% -> %"T"%) >]).
2024-07-27 13:30:12 +02:00
apply T_TypeAbs.
apply T_Abs.
2024-08-21 20:04:20 +02:00
apply T_Var.
apply C_take.
2024-08-21 20:04:20 +02:00
apply TSubRepr_Refl.
apply TEq_Alpha.
apply TAlpha_Rename.
Qed.
Example typing3 :
ctx_empty |- [{
Λ"T" Λ"U" λ"x",%"T"% λ"y",%"U"% %"y"%
}] \is [<
"S","T",(%"S"%->%"T"%->%"T"%)
>].
2024-08-21 20:04:20 +02:00
Proof.
intros.
apply T_DescendImplicit with (τ:=[< "T","U",(%"T"%->%"U"%->%"U"%) >]) (τ':=[< "S","T",(%"S"%->%"T"%->%"T"%) >]).
2024-08-21 20:04:20 +02:00
apply T_TypeAbs, T_TypeAbs, T_Abs.
apply T_Abs.
apply T_Var.
apply C_take.
2024-08-21 20:04:20 +02:00
apply TSubRepr_Refl.
apply TEq_Trans with (y:= [< "S","U",(%"S"%->%"U"%->%"U"%) >] ).
2024-08-21 20:04:20 +02:00
apply TEq_Alpha.
apply TAlpha_Rename.
apply TEq_Alpha.
apply TAlpha_SubUniv.
apply TAlpha_Rename.
Qed.
2024-08-22 08:31:14 +02:00
Example typing4 : (is_well_typed
[{ Λ"T" Λ"U" λ"x",%"T"% λ"y",%"U"% %"x"% }]
2024-08-22 08:31:14 +02:00
).
Proof.
unfold is_well_typed.
exists [< "T","U",%"T"%->%"U"%->%"T"% >].
2024-09-04 12:39:15 +02:00
apply T_TypeAbs.
apply T_TypeAbs.
apply T_Abs.
apply T_Abs.
apply T_Var.
apply C_shuffle, C_take.
Qed.
Example typing5 :
(ctx_assign "60" [< $""$ >]
(ctx_assign "360" [< $""$ >]
(ctx_assign "/" [< $""$ -> $""$ -> $""$ >]
ctx_empty)))
|-
[{
let "deg2turns" :=
(λ"x",$"Angle"$~$"Degrees"$~$""$
morph (((%"/"% %"x"%) %"360"%) as $"Angle"$~$"Turns"$))
in ( %"deg2turns"% (%"60"% as $"Angle"$~$"Degrees"$) )
}]
\is
[<
$"Angle"$~$"Turns"$~$""$
>]
.
Proof.
apply T_Let with (σ:=[< $"Angle"$~$"Degrees"$~$""$ ->morph $"Angle"$~$"Turns"$~$""$ >]).
apply T_MorphAbs.
apply T_DescendImplicit with (τ:=(type_ladder [<$"Angle"$~$"Turns"$>] [<$""$>])).
2: apply TSubRepr_Refl, TEq_LadderAssocLR.
apply T_Ascend with (τ:=[<$""$>]) (τ':=[<$"Angle"$~$"Turns"$>]).
apply T_App with (σ := [< $""$ >]) (σ' := [< $""$ >]).
apply T_App with (σ := [< $""$ >]) (σ' := [< $""$ >]).
apply T_Var.
apply C_shuffle. apply C_shuffle. apply C_shuffle. apply C_take.
apply T_DescendImplicit with (τ := [< $"Angle"$~$"Degrees"$~$""$ >]).
apply T_Var.
apply C_take.
apply TSubRepr_Ladder. apply TSubRepr_Ladder.
apply TSubRepr_Refl. apply TEq_Refl.
apply M_Sub.
apply TSubRepr_Refl. apply TEq_Refl.
apply T_Var.
apply C_shuffle, C_shuffle, C_take.
apply M_Sub.
apply TSubRepr_Refl.
apply TEq_Refl.
apply T_App with (σ:=[<$"Angle"$~$"Degrees"$~$""$>]) (σ':=[<$"Angle"$~$"Degrees"$~$""$>]).
apply T_MorphFun.
2024-09-04 12:39:15 +02:00
apply T_Var.
2024-08-22 08:31:14 +02:00
apply C_take.
apply T_DescendImplicit with (τ:=(type_ladder [<$"Angle"$~$"Degrees"$>] [<$""$>])).
2: apply TSubRepr_Refl, TEq_LadderAssocLR.
apply T_Ascend.
apply T_Var.
apply C_shuffle. apply C_take.
apply M_Sub. apply TSubRepr_Refl. apply TEq_Refl.
2024-08-22 08:31:14 +02:00
Qed.
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.