add notation for sequence types & use notations everywhere
This commit is contained in:
parent
633843d9c7
commit
6e5c832db7
7 changed files with 260 additions and 226 deletions
148
coq/bbencode.v
148
coq/bbencode.v
|
@ -4,77 +4,101 @@ Require Import terms.
|
|||
Require Import subst.
|
||||
Require Import smallstep.
|
||||
|
||||
Include Terms.
|
||||
Include Subst.
|
||||
Include Smallstep.
|
||||
|
||||
Open Scope ladder_type_scope.
|
||||
Open Scope ladder_expr_scope.
|
||||
|
||||
(* let bb_zero = Λα ↦ λs: α->α ↦ λz: α ↦ z
|
||||
* ∀α.(α->α)->α->α
|
||||
*)
|
||||
Definition bb_zero : expr_term :=
|
||||
(expr_ty_abs "α"
|
||||
(expr_abs "s" (type_fun (type_var "α") (type_var "α"))
|
||||
(expr_abs "z" (type_var "α")
|
||||
(expr_var "z")))).
|
||||
|
||||
(* let bb_one = Λα ↦ λs: α->α ↦ λz: α ↦ s z
|
||||
*)
|
||||
Definition bb_one : expr_term :=
|
||||
(expr_ty_abs "α"
|
||||
(expr_abs "s" (type_fun (type_var "α") (type_var "α"))
|
||||
(expr_abs "z" (type_var "α")
|
||||
(expr_app (expr_var "s") (expr_var "z"))))).
|
||||
Definition bb_zero : expr_term := [{
|
||||
(Λ"α" ↦
|
||||
λ"s",(%"α"% -> %"α"%) ↦
|
||||
λ"z",(%"α"%) ↦
|
||||
%"z"%)
|
||||
as $"ℕ"$~$"BBNat"$
|
||||
}].
|
||||
|
||||
(* let bb_two = Λα ↦ λs: α->α ↦ λz: α ↦ s (s z)
|
||||
*)
|
||||
Definition bb_two : expr_term :=
|
||||
(expr_ty_abs "α"
|
||||
(expr_abs "s" (type_fun (type_var "α") (type_var "α"))
|
||||
(expr_abs "z" (type_var "α")
|
||||
(expr_app (expr_var "s") (expr_app (expr_var "s") (expr_var "z")))))).
|
||||
Definition bb_one : expr_term := [{
|
||||
(Λ"α" ↦
|
||||
λ"s",(%"α"% -> %"α"%) ↦
|
||||
λ"z",(%"α"%) ↦
|
||||
(%"s"% %"z"%))
|
||||
as $"ℕ"$~$"BBNat"$
|
||||
}].
|
||||
|
||||
Definition bb_two : expr_term := [{
|
||||
(Λ"α" ↦
|
||||
λ"s",(%"α"% -> %"α"%) ↦
|
||||
λ"z",(%"α"%) ↦
|
||||
(%"s"% (%"s"% %"z"%)))
|
||||
as $"ℕ"$~$"BBNat"$
|
||||
}].
|
||||
|
||||
Definition bb_succ : expr_term :=
|
||||
(expr_abs "n" (type_ladder (type_id "ℕ")
|
||||
(type_ladder (type_id "BBNat")
|
||||
(type_univ "α"
|
||||
(type_fun (type_fun (type_var "α") (type_var "α"))
|
||||
(type_fun (type_var "α") (type_var "α"))))))
|
||||
Definition bb_succ : expr_term := [{
|
||||
λ"n",$"ℕ"$
|
||||
~$"BBNat"$
|
||||
~(∀"β",(%"β'"%->%"β"%)->%"β"%->%"β"%)
|
||||
↦
|
||||
((Λ"α" ↦
|
||||
λ"s",((%"α"%->%"α"%)->%"α"%->%"α"%) ↦
|
||||
λ"z",%"α"% ↦
|
||||
(%"s"% ((%"n"% # %"α"%) %"s"% %"z"%)))
|
||||
as $"ℕ"$~$"BBNat"$)
|
||||
}].
|
||||
|
||||
(expr_ascend (type_ladder (type_id "ℕ") (type_id "BBNat"))
|
||||
(expr_ty_abs "α"
|
||||
(expr_abs "s" (type_fun (type_var "α") (type_var "α"))
|
||||
(expr_abs "z" (type_var "α")
|
||||
(expr_app (expr_var "s")
|
||||
(expr_app (expr_app
|
||||
(expr_ty_app (expr_var "n") (type_var "α"))
|
||||
(expr_var "s"))
|
||||
(expr_var "z")))))))).
|
||||
|
||||
Definition e1 : expr_term :=
|
||||
(expr_let "bb-zero" bb_zero
|
||||
(expr_app (expr_app (expr_var "+") (expr_var "bb-zero")) (expr_var "bb-zero"))
|
||||
).
|
||||
|
||||
Definition t1 : expr_term := (expr_app (expr_var "x") (expr_var "x")).
|
||||
|
||||
Compute (expr_subst "x"
|
||||
(expr_ty_abs "α" (expr_abs "a" (type_var "α") (expr_var "a")))
|
||||
bb_one
|
||||
).
|
||||
|
||||
Example example_let_reduction :
|
||||
e1 -->β (expr_app (expr_app (expr_var "+") bb_zero) bb_zero).
|
||||
Example example_reduction :
|
||||
(expr_app bb_succ bb_one) -->β* bb_two
|
||||
.
|
||||
Proof.
|
||||
apply E_Let.
|
||||
|
||||
apply Multi_Step with (y:=[{
|
||||
(λ"n",$"ℕ"$
|
||||
~$"BBNat"$
|
||||
~(∀"β",(%"β'"%->%"β"%)->%"β"%->%"β"%)
|
||||
↦
|
||||
((Λ"α" ↦
|
||||
λ"s",((%"α"%->%"α"%)->%"α"%->%"α"%) ↦
|
||||
λ"z",%"α"% ↦
|
||||
(%"s"% ((%"n"% # %"α"%) %"s"% %"z"%)))
|
||||
as $"ℕ"$~$"BBNat"$))
|
||||
|
||||
(Λ"β" ↦
|
||||
λ"s",(%"β"% -> %"β"%) ↦
|
||||
λ"z",(%"β"%) ↦
|
||||
(%"s"% %"z"%))
|
||||
as $"ℕ"$~$"BBNat"$
|
||||
}]).
|
||||
Admitted.
|
||||
(*
|
||||
apply E_Alpha.
|
||||
apply E_App2.
|
||||
apply EAlpha_Rename.
|
||||
|
||||
apply Multi_Step with (y:=[{
|
||||
(Λ"α" ↦
|
||||
λ"s",((%"α"%->%"α"%)->%"α"%->%"α"%) ↦
|
||||
λ"z",%"α"% ↦
|
||||
(%"s"% (
|
||||
(((Λ"β" ↦
|
||||
λ"s",(%"β"% -> %"β"%) ↦
|
||||
λ"z",(%"β"%) ↦
|
||||
(%"s"% %"z"%)) as $"ℕ"$~$"BBNat"$)
|
||||
# %"α"%) %"s"% %"z"%)))
|
||||
as $"ℕ"$~$"BBNat"$ }]).
|
||||
|
||||
apply E_AppLam.
|
||||
|
||||
apply Multi_Step with (y:=[{
|
||||
(Λ"α" ↦
|
||||
λ"s",((%"α"%->%"α"%)->%"α"%->%"α"%) ↦
|
||||
λ"z",%"α"% ↦
|
||||
(%"s"% (
|
||||
(((Λ"α'" ↦
|
||||
λ"s",(%"α'"% -> %"α'"%) ↦
|
||||
λ"z",(%"α'"%) ↦
|
||||
(%"s"% %"z"%)) as $"ℕ"$~$"BBNat"$)
|
||||
# %"α"%) %"s"% %"z"%)))
|
||||
as $"ℕ"$~$"BBNat"$ }]).
|
||||
}].
|
||||
Qed.
|
||||
*)
|
||||
|
||||
Compute (expr_app bb_succ bb_zero) -->β bb_one.
|
||||
|
||||
Example example_succ :
|
||||
(expr_app bb_succ bb_zero) -->β bb_one.
|
||||
Proof.
|
||||
Admitted.
|
||||
|
|
102
coq/equiv.v
102
coq/equiv.v
|
@ -38,46 +38,44 @@ From Coq Require Import Strings.String.
|
|||
|
||||
Reserved Notation "S '--->α' T" (at level 40).
|
||||
Inductive type_conv_alpha : type_term -> type_term -> Prop :=
|
||||
| TAlpha_Rename : forall x y t t',
|
||||
|
||||
(type_subst1 x (type_var y) t t') ->
|
||||
(type_univ x t) --->α (type_univ y t')
|
||||
| TAlpha_Rename : forall x y t,
|
||||
(type_univ x t) --->α (type_univ y (type_subst x (type_var y) t))
|
||||
|
||||
| TAlpha_SubUniv : forall x τ τ',
|
||||
(τ --->α τ') ->
|
||||
(type_univ x τ) --->α (type_univ x τ')
|
||||
[< ∀x,τ >] --->α [< ∀x,τ' >]
|
||||
|
||||
| TAlpha_SubSpec1 : forall τ1 τ1' τ2,
|
||||
(τ1 --->α τ1') ->
|
||||
(type_spec τ1 τ2) --->α (type_spec τ1' τ2)
|
||||
[< <τ1 τ2> >] --->α [< <τ1' τ2> >]
|
||||
|
||||
| TAlpha_SubSpec2 : forall τ1 τ2 τ2',
|
||||
(τ2 --->α τ2') ->
|
||||
(type_spec τ1 τ2) --->α (type_spec τ1 τ2')
|
||||
[< <τ1 τ2> >] --->α [< <τ1 τ2'> >]
|
||||
|
||||
| TAlpha_SubFun1 : forall τ1 τ1' τ2,
|
||||
(τ1 --->α τ1') ->
|
||||
(type_fun τ1 τ2) --->α (type_fun τ1' τ2)
|
||||
[< τ1 -> τ2 >] --->α [< τ1' -> τ2 >]
|
||||
|
||||
| TAlpha_SubFun2 : forall τ1 τ2 τ2',
|
||||
(τ2 --->α τ2') ->
|
||||
(type_fun τ1 τ2) --->α (type_fun τ1 τ2')
|
||||
[< τ1 -> τ2 >] --->α [< τ1 -> τ2' >]
|
||||
|
||||
| TAlpha_SubMorph1 : forall τ1 τ1' τ2,
|
||||
(τ1 --->α τ1') ->
|
||||
(type_morph τ1 τ2) --->α (type_morph τ1' τ2)
|
||||
[< τ1 ->morph τ2 >] --->α [< τ1' ->morph τ2 >]
|
||||
|
||||
| TAlpha_SubMorph2 : forall τ1 τ2 τ2',
|
||||
(τ2 --->α τ2') ->
|
||||
(type_morph τ1 τ2) --->α (type_morph τ1 τ2')
|
||||
[< τ1 ->morph τ2 >] --->α [< τ1 ->morph τ2' >]
|
||||
|
||||
| TAlpha_SubLadder1 : forall τ1 τ1' τ2,
|
||||
(τ1 --->α τ1') ->
|
||||
(type_ladder τ1 τ2) --->α (type_ladder τ1' τ2)
|
||||
[< τ1 ~ τ2 >] --->α [< τ1' ~ τ2 >]
|
||||
|
||||
| TAlpha_SubLadder2 : forall τ1 τ2 τ2',
|
||||
(τ2 --->α τ2') ->
|
||||
(type_ladder τ1 τ2) --->α (type_ladder τ1 τ2')
|
||||
[< τ1 ~ τ2 >] --->α [< τ1 ~ τ2' >]
|
||||
|
||||
where "S '--->α' T" := (type_conv_alpha S T).
|
||||
|
||||
|
@ -86,38 +84,48 @@ Lemma type_alpha_symm :
|
|||
forall σ τ,
|
||||
(σ --->α τ) -> (τ --->α σ).
|
||||
Proof.
|
||||
(* TODO *)
|
||||
intros.
|
||||
induction H.
|
||||
- admit.
|
||||
- auto using TAlpha_Rename,type_subst,type_subst_symm, TAlpha_SubUniv.
|
||||
- auto using TAlpha_SubSpec1.
|
||||
- auto using TAlpha_SubSpec2.
|
||||
- auto using TAlpha_SubFun1.
|
||||
- auto using TAlpha_SubFun2.
|
||||
- auto using TAlpha_SubMorph1.
|
||||
- auto using TAlpha_SubMorph2.
|
||||
- auto using TAlpha_SubLadder1.
|
||||
- auto using TAlpha_SubLadder2.
|
||||
Admitted.
|
||||
|
||||
|
||||
(** Define all rewrite steps $\label{coq:type-dist}$ *)
|
||||
|
||||
Reserved Notation "S '-->distribute-ladder' T" (at level 40).
|
||||
Inductive type_distribute_ladder : type_term -> type_term -> Prop :=
|
||||
| L_DistributeOverSpec1 : forall x x' y,
|
||||
(type_spec (type_ladder x x') y)
|
||||
[< <x~x' y> >]
|
||||
-->distribute-ladder
|
||||
(type_ladder (type_spec x y) (type_spec x' y))
|
||||
[< <x y>~<x' y> >]
|
||||
|
||||
| L_DistributeOverSpec2 : forall x y y',
|
||||
(type_spec x (type_ladder y y'))
|
||||
[< <x y~y'> >]
|
||||
-->distribute-ladder
|
||||
(type_ladder (type_spec x y) (type_spec x y'))
|
||||
[< <x y>~<x y'> >]
|
||||
|
||||
| L_DistributeOverFun1 : forall x x' y,
|
||||
(type_fun (type_ladder x x') y)
|
||||
[< (x~x' -> y) >]
|
||||
-->distribute-ladder
|
||||
(type_ladder (type_fun x y) (type_fun x' y))
|
||||
[< (x -> y) ~ (x' -> y) >]
|
||||
|
||||
| L_DistributeOverFun2 : forall x y y',
|
||||
(type_fun x (type_ladder y y'))
|
||||
[< (x -> y~y') >]
|
||||
-->distribute-ladder
|
||||
(type_ladder (type_fun x y) (type_fun x y'))
|
||||
[< (x -> y) ~ (x -> y') >]
|
||||
|
||||
| L_DistributeOverMorph1 : forall x x' y,
|
||||
(type_morph (type_ladder x x') y)
|
||||
[< (x~x' ->morph y) >]
|
||||
-->distribute-ladder
|
||||
(type_ladder (type_morph x y) (type_morph x' y))
|
||||
[< (x ->morph y) ~ (x' ->morph y) >]
|
||||
|
||||
| L_DistributeOverMorph2 : forall x y y',
|
||||
(type_morph x (type_ladder y y'))
|
||||
|
@ -131,34 +139,34 @@ where "S '-->distribute-ladder' T" := (type_distribute_ladder S T).
|
|||
Reserved Notation "S '-->condense-ladder' T" (at level 40).
|
||||
Inductive type_condense_ladder : type_term -> type_term -> Prop :=
|
||||
| L_CondenseOverSpec1 : forall x x' y,
|
||||
(type_ladder (type_spec x y) (type_spec x' y))
|
||||
[< <x y>~<x' y> >]
|
||||
-->condense-ladder
|
||||
(type_spec (type_ladder x x') y)
|
||||
[< <x~x' y> >]
|
||||
|
||||
| L_CondenseOverSpec2 : forall x y y',
|
||||
(type_ladder (type_spec x y) (type_spec x y'))
|
||||
[< <x y>~<x y'> >]
|
||||
-->condense-ladder
|
||||
(type_spec x (type_ladder y y'))
|
||||
[< <x y~y'> >]
|
||||
|
||||
| L_CondenseOverFun1 : forall x x' y,
|
||||
(type_ladder (type_fun x y) (type_fun x' y))
|
||||
[< (x -> y) ~ (x' -> y) >]
|
||||
-->condense-ladder
|
||||
(type_fun (type_ladder x x') y)
|
||||
[< (x~x') -> y >]
|
||||
|
||||
| L_CondenseOverFun2 : forall x y y',
|
||||
(type_ladder (type_fun x y) (type_fun x y'))
|
||||
[< (x -> y) ~ (x -> y') >]
|
||||
-->condense-ladder
|
||||
(type_fun x (type_ladder y y'))
|
||||
[< (x -> y~y') >]
|
||||
|
||||
| L_CondenseOverMorph1 : forall x x' y,
|
||||
(type_ladder (type_morph x y) (type_morph x' y))
|
||||
[< (x ->morph y) ~ (x' ->morph y) >]
|
||||
-->condense-ladder
|
||||
(type_morph (type_ladder x x') y)
|
||||
[< (x~x' ->morph y) >]
|
||||
|
||||
| L_CondenseOverMorph2 : forall x y y',
|
||||
(type_ladder (type_morph x y) (type_morph x y'))
|
||||
[< (x ->morph y) ~ (x ->morph y') >]
|
||||
-->condense-ladder
|
||||
(type_morph x (type_ladder y y'))
|
||||
[< (x ->morph y~y') >]
|
||||
|
||||
where "S '-->condense-ladder' T" := (type_condense_ladder S T).
|
||||
|
||||
|
@ -213,14 +221,14 @@ Inductive type_eq : type_term -> type_term -> Prop :=
|
|||
x === z
|
||||
|
||||
| TEq_LadderAssocLR : forall x y z,
|
||||
(type_ladder (type_ladder x y) z)
|
||||
[< (x~y)~z >]
|
||||
===
|
||||
(type_ladder x (type_ladder y z))
|
||||
[< x~(y~z) >]
|
||||
|
||||
| TEq_LadderAssocRL : forall x y z,
|
||||
(type_ladder x (type_ladder y z))
|
||||
[< x~(y~z) >]
|
||||
===
|
||||
(type_ladder (type_ladder x y) z)
|
||||
[< (x~y)~z >]
|
||||
|
||||
| TEq_Alpha : forall x y,
|
||||
x --->α y ->
|
||||
|
@ -378,7 +386,7 @@ Admitted.
|
|||
*)
|
||||
|
||||
Example example_flat_type :
|
||||
(type_is_flat (type_spec (type_id "PosInt") (type_id "10"))).
|
||||
type_is_flat [< <$"PosInt"$ $"10"$> >].
|
||||
Proof.
|
||||
apply FlatApp.
|
||||
apply FlatId.
|
||||
|
@ -386,17 +394,15 @@ Proof.
|
|||
Qed.
|
||||
|
||||
Example example_lnf_type :
|
||||
(type_is_lnf
|
||||
(type_ladder (type_spec (type_id "Seq") (type_id "Char"))
|
||||
(type_spec (type_id "Seq") (type_id "Byte")))).
|
||||
type_is_lnf [< [$"Char"$] ~ [$"Byte"$] >].
|
||||
Proof.
|
||||
Admitted.
|
||||
|
||||
Example example_type_eq :
|
||||
(type_spec (type_id "Seq") (type_ladder (type_id "Char") (type_id "Byte")))
|
||||
[< [$"Char"$ ~ $"Byte"$] >]
|
||||
===
|
||||
(type_ladder (type_spec (type_id "Seq") (type_id "Char"))
|
||||
(type_spec (type_id "Seq") (type_id "Byte"))).
|
||||
[< [$"Char"$] ~ [$"Byte"$] >]
|
||||
.
|
||||
Proof.
|
||||
apply TEq_Distribute.
|
||||
apply L_DistributeOverSpec2.
|
||||
|
|
45
coq/morph.v
45
coq/morph.v
|
@ -8,13 +8,15 @@ Require Import context.
|
|||
(* Given a context, there is a morphism path from τ to τ' *)
|
||||
Reserved Notation "Γ '|-' σ '~>' τ" (at level 101, σ at next level, τ at next level).
|
||||
|
||||
Open Scope ladder_expr_scope.
|
||||
|
||||
Inductive morphism_path : context -> type_term -> type_term -> Prop :=
|
||||
| M_Sub : forall Γ τ τ',
|
||||
(τ :<= τ') ->
|
||||
(Γ |- τ ~> τ')
|
||||
|
||||
| M_Single : forall Γ h τ τ',
|
||||
(context_contains Γ h (type_morph τ τ')) ->
|
||||
(context_contains Γ h [< τ ->morph τ' >]) ->
|
||||
(Γ |- τ ~> τ')
|
||||
|
||||
| M_Chain : forall Γ τ τ' τ'',
|
||||
|
@ -24,58 +26,57 @@ Inductive morphism_path : context -> type_term -> type_term -> Prop :=
|
|||
|
||||
| M_Lift : forall Γ σ τ τ',
|
||||
(Γ |- τ ~> τ') ->
|
||||
(Γ |- (type_ladder σ τ) ~> (type_ladder σ τ'))
|
||||
(Γ |- [< σ ~ τ >] ~> [< σ ~ τ' >])
|
||||
|
||||
| M_MapSeq : forall Γ τ τ',
|
||||
(Γ |- τ ~> τ') ->
|
||||
(Γ |- (type_spec (type_id "Seq") τ) ~> (type_spec (type_id "Seq") τ'))
|
||||
(Γ |- [< [τ] >] ~> [< [τ'] >])
|
||||
|
||||
where "Γ '|-' s '~>' t" := (morphism_path Γ s t).
|
||||
|
||||
Lemma id_morphism_path : forall Γ τ, Γ |- τ ~> τ.
|
||||
Proof.
|
||||
intros.
|
||||
apply M_Sub, TSubRepr_Refl, TEq_Refl.
|
||||
Qed.
|
||||
|
||||
Inductive translate_morphism_path : context -> type_term -> type_term -> expr_term -> Prop :=
|
||||
| Translate_Subtype : forall Γ τ τ',
|
||||
| Translate_Descend : forall Γ τ τ',
|
||||
(τ :<= τ') ->
|
||||
(translate_morphism_path Γ τ τ'
|
||||
(expr_morph "x" τ (expr_var "x")))
|
||||
(expr_morph "x" τ [{ %"x"% des τ' }]))
|
||||
|
||||
| Translate_Lift : forall Γ σ τ τ' m,
|
||||
(Γ |- τ ~> τ') ->
|
||||
(translate_morphism_path Γ τ τ' m) ->
|
||||
(translate_morphism_path Γ (type_ladder σ τ) (type_ladder σ τ')
|
||||
(expr_morph "x" (type_ladder σ τ)
|
||||
(expr_ascend σ (expr_app m (expr_descend τ (expr_var "x"))))))
|
||||
(translate_morphism_path Γ [< σ ~ τ >] [< σ ~ τ' >]
|
||||
(expr_morph "x" [< σ ~ τ >] [{ (m (%"x"% des τ)) as σ }]))
|
||||
|
||||
| Translate_Single : forall Γ h τ τ',
|
||||
(context_contains Γ h (type_morph τ τ')) ->
|
||||
(translate_morphism_path Γ τ τ' (expr_var h))
|
||||
(context_contains Γ h [< τ ->morph τ' >]) ->
|
||||
(translate_morphism_path Γ τ τ' [{ %h% }])
|
||||
|
||||
| Translate_Chain : forall Γ τ τ' τ'' m1 m2,
|
||||
(translate_morphism_path Γ τ τ' m1) ->
|
||||
(translate_morphism_path Γ τ' τ'' m2) ->
|
||||
(translate_morphism_path Γ τ τ''
|
||||
(expr_morph "x" τ (expr_app m2 (expr_app m1 (expr_var "x")))))
|
||||
(expr_morph "x" τ [{ m2 (m1 %"x"%) }]))
|
||||
|
||||
| Translate_MapSeq : forall Γ τ τ' m,
|
||||
(translate_morphism_path Γ τ τ' m) ->
|
||||
(translate_morphism_path Γ
|
||||
(type_spec (type_id "Seq") τ)
|
||||
(type_spec (type_id "Seq") τ')
|
||||
(expr_morph "xs"
|
||||
(type_spec (type_id "Seq") τ)
|
||||
(expr_app (expr_app (expr_ty_app (expr_ty_app
|
||||
(expr_var "map") τ) τ') m)
|
||||
(expr_var "xs"))))
|
||||
(translate_morphism_path Γ [< [τ] >] [< [τ'] >]
|
||||
[{
|
||||
λ"xs",[τ] ↦morph (%"map"% # τ # τ' m %"xs"%)
|
||||
}])
|
||||
.
|
||||
|
||||
|
||||
Example morphism_paths :
|
||||
(ctx_assign "degrees-to-turns" [< $"Angle"$~$"Degrees"$~$"ℝ"$ ->morph $"Angle"$~$"Turns"$~$"ℝ"$ >]
|
||||
(ctx_assign "turns-to-radians" [< $"Angle"$~$"Turns"$~$"ℝ"$ ->morph $"Angle"$~$"Radians"$~$"ℝ"$ >]
|
||||
ctx_empty))
|
||||
|
||||
|- [< <$"Seq"$ $"Hue"$~$"Angle"$~$"Degrees"$~$"ℝ"$> >]
|
||||
~> [< <$"Seq"$ $"Hue"$~$"Angle"$~$"Radians"$~$"ℝ"$> >]
|
||||
|- [< [ $"Hue"$~$"Angle"$~$"Degrees"$~$"ℝ"$ ] >]
|
||||
~> [< [ $"Hue"$~$"Angle"$~$"Radians"$~$"ℝ"$ ] >]
|
||||
.
|
||||
Proof.
|
||||
intros.
|
||||
|
|
|
@ -16,19 +16,19 @@ Inductive expr_alpha : expr_term -> expr_term -> Prop :=
|
|||
|
||||
| EAlpha_SubAbs : forall x τ e e',
|
||||
(e -->α e') ->
|
||||
(expr_abs x τ e) -->α (expr_abs x τ e')
|
||||
[{ λ x , τ ↦ e }] -->α [{ λ x , τ ↦ e' }]
|
||||
|
||||
| EAlpha_SubTyAbs : forall α e e',
|
||||
(e -->α e') ->
|
||||
(expr_ty_abs α e) -->α (expr_ty_abs α e')
|
||||
[{ Λ α ↦ e }] -->α [{ Λ α ↦ e' }]
|
||||
|
||||
| EAlpha_SubApp1 : forall e1 e1' e2,
|
||||
(e1 -->α e1') ->
|
||||
(expr_app e1 e2) -->α (expr_app e1' e2)
|
||||
[{ e1 e2 }] -->α [{ e1' e2 }]
|
||||
|
||||
| EAlpha_SubApp2 : forall e1 e2 e2',
|
||||
(e2 -->α e2') ->
|
||||
(expr_app e1 e2) -->α (expr_app e1 e2')
|
||||
[{ e1 e2 }] -->α [{ e1 e2' }]
|
||||
|
||||
where "s '-->α' t" := (expr_alpha s t).
|
||||
|
||||
|
@ -45,45 +45,45 @@ Qed.
|
|||
Inductive beta_step : expr_term -> expr_term -> Prop :=
|
||||
| E_App1 : forall e1 e1' e2,
|
||||
e1 -->β e1' ->
|
||||
(expr_app e1 e2) -->β (expr_app e1' e2)
|
||||
[{ e1 e2 }] -->β [{ e1' e2 }]
|
||||
|
||||
| E_App2 : forall v1 e2 e2',
|
||||
(is_value v1) ->
|
||||
e2 -->β e2' ->
|
||||
(expr_app v1 e2) -->β (expr_app v1 e2')
|
||||
[{ v1 e2 }] -->β [{ v1 e2' }]
|
||||
|
||||
| E_TypApp : forall e e' τ,
|
||||
e -->β e' ->
|
||||
(expr_ty_app e τ) -->β (expr_ty_app e' τ)
|
||||
[{ Λ τ ↦ e }] -->β [{ Λ τ ↦ e' }]
|
||||
|
||||
| E_TypAppLam : forall α e τ,
|
||||
(expr_ty_app (expr_ty_abs α e) τ) -->β (expr_specialize α τ e)
|
||||
[{ (Λ α ↦ e) # τ }] -->β (expr_specialize α τ e)
|
||||
|
||||
| E_AppLam : forall x τ e a,
|
||||
(expr_app (expr_abs x τ e) a) -->β (expr_subst x a e)
|
||||
[{ (λ x , τ ↦ e) a }] -->β (expr_subst x a e)
|
||||
|
||||
| E_AppMorph : forall x τ e a,
|
||||
(expr_app (expr_morph x τ e) a) -->β (expr_subst x a e)
|
||||
[{ (λ x , τ ↦morph e) a }] -->β (expr_subst x a e)
|
||||
|
||||
| E_Let : forall x e a,
|
||||
(expr_let x a e) -->β (expr_subst x a e)
|
||||
[{ let x := a in e }] -->β (expr_subst x a e)
|
||||
|
||||
| E_StripAscend : forall τ e,
|
||||
(expr_ascend τ e) -->β e
|
||||
[{ e as τ }] -->β e
|
||||
|
||||
| E_StripDescend : forall τ e,
|
||||
(expr_descend τ e) -->β e
|
||||
[{ e des τ }] -->β e
|
||||
|
||||
| E_Ascend : forall τ e e',
|
||||
(e -->β e') ->
|
||||
(expr_ascend τ e) -->β (expr_ascend τ e')
|
||||
[{ e as τ }] -->β [{ e' as τ }]
|
||||
|
||||
| E_AscendCollapse : forall τ' τ e,
|
||||
(expr_ascend τ' (expr_ascend τ e)) -->β (expr_ascend (type_ladder τ' τ) e)
|
||||
[{ (e as τ) as τ' }] -->β [{ e as (τ'~τ) }]
|
||||
|
||||
| E_DescendCollapse : forall τ' τ e,
|
||||
(τ':<=τ) ->
|
||||
(expr_descend τ (expr_descend τ' e)) -->β (expr_descend τ e)
|
||||
[{ (e des τ') des τ }] -->β [{ e des τ }]
|
||||
|
||||
where "s '-->β' t" := (beta_step s t).
|
||||
|
||||
|
@ -102,7 +102,7 @@ Notation " s -->β* t " := (multi beta_step s t) (at level 40).
|
|||
Example reduce1 :
|
||||
[{
|
||||
let "deg2turns" :=
|
||||
(λ"x" $"Angle"$~$"Degrees"$~$"ℝ"$
|
||||
(λ"x",$"Angle"$~$"Degrees"$~$"ℝ"$
|
||||
↦morph ((%"/"% (%"x"% des $"ℝ"$) %"360"%) as $"Angle"$~$"Turns"$))
|
||||
in ( %"deg2turns"% (%"60"% as $"Angle"$~$"Degrees"$) )
|
||||
}]
|
||||
|
@ -111,7 +111,7 @@ Example reduce1 :
|
|||
((%"/"% %"60"%) %"360"%) as $"Angle"$~$"Turns"$
|
||||
}].
|
||||
Proof.
|
||||
apply Multi_Step with (y:=[{ (λ"x" $"Angle"$~$"Degrees"$~$"ℝ"$
|
||||
apply Multi_Step with (y:=[{ (λ"x",$"Angle"$~$"Degrees"$~$"ℝ"$
|
||||
↦morph (((%"/"% (%"x"% des $"ℝ"$)) %"360"%) as $"Angle"$~$"Turns"$)) (%"60"% as $"Angle"$~$"Degrees"$) }]).
|
||||
apply E_Let.
|
||||
|
||||
|
|
|
@ -8,7 +8,6 @@ Require Import morph.
|
|||
Require Import smallstep.
|
||||
Require Import typing.
|
||||
|
||||
|
||||
Lemma typing_weakening : forall Γ e τ x σ,
|
||||
(Γ |- e \is τ) ->
|
||||
((ctx_assign x σ Γ) |- e \is τ)
|
||||
|
|
44
coq/terms.v
44
coq/terms.v
|
@ -64,58 +64,66 @@ Declare Custom Entry ladder_type.
|
|||
Declare Custom Entry ladder_expr.
|
||||
|
||||
Notation "[< t >]" := t
|
||||
(t custom ladder_type at level 80) : ladder_type_scope.
|
||||
(t custom ladder_type at level 99) : ladder_type_scope.
|
||||
Notation "t" := t
|
||||
(in custom ladder_type at level 0, t ident) : ladder_type_scope.
|
||||
Notation "'∀' x ',' t" := (type_univ x t)
|
||||
(t custom ladder_type at level 80, in custom ladder_type at level 80, x constr).
|
||||
Notation "'<' σ τ '>'" := (type_spec σ τ)
|
||||
(in custom ladder_type at level 80, left associativity) : ladder_type_scope.
|
||||
Notation "'(' τ ')'" := τ
|
||||
Notation "'[' τ ']'" := (type_spec (type_id "Seq") τ)
|
||||
(in custom ladder_type at level 70) : ladder_type_scope.
|
||||
Notation "'(' τ ')'" := τ
|
||||
(in custom ladder_type at level 5) : ladder_type_scope.
|
||||
Notation "σ '->' τ" := (type_fun σ τ)
|
||||
(in custom ladder_type at level 75, right associativity) : ladder_type_scope.
|
||||
Notation "σ '->morph' τ" := (type_morph σ τ)
|
||||
(in custom ladder_type at level 75, right associativity, τ at level 80) : ladder_type_scope.
|
||||
Notation "σ '~' τ" := (type_ladder σ τ)
|
||||
(in custom ladder_type at level 70, right associativity) : ladder_type_scope.
|
||||
(in custom ladder_type at level 20, right associativity) : ladder_type_scope.
|
||||
Notation "'$' x '$'" := (type_id x%string)
|
||||
(in custom ladder_type at level 0, x constr) : ladder_type_scope.
|
||||
(in custom ladder_type at level 10, x constr) : ladder_type_scope.
|
||||
Notation "'%' x '%'" := (type_var x%string)
|
||||
(in custom ladder_type at level 0, x constr) : ladder_type_scope.
|
||||
(in custom ladder_type at level 10, x constr) : ladder_type_scope.
|
||||
|
||||
Notation "[{ e }]" := e
|
||||
(e custom ladder_expr at level 80) : ladder_expr_scope.
|
||||
(e custom ladder_expr at level 99) : ladder_expr_scope.
|
||||
Notation "e" := e
|
||||
(in custom ladder_expr at level 0, e ident) : ladder_expr_scope.
|
||||
Notation "'%' x '%'" := (expr_var x%string)
|
||||
(in custom ladder_expr at level 0, x constr) : ladder_expr_scope.
|
||||
(in custom ladder_expr at level 10, x constr) : ladder_expr_scope.
|
||||
Notation "'Λ' t '↦' e" := (expr_ty_abs t e)
|
||||
(in custom ladder_expr at level 10, t constr, e custom ladder_expr at level 80) : ladder_expr_scope.
|
||||
Notation "'λ' x τ '↦' e" := (expr_abs x τ e)
|
||||
(in custom ladder_expr at level 10, x constr, τ custom ladder_type at level 99, e custom ladder_expr at level 99) :ladder_expr_scope.
|
||||
Notation "'λ' x τ '↦morph' e" := (expr_morph x τ e)
|
||||
(in custom ladder_expr at level 10, x constr, τ custom ladder_type at level 99, e custom ladder_expr at level 99) :ladder_expr_scope.
|
||||
(in custom ladder_expr at level 10, t constr, e custom ladder_expr at level 80, right associativity) : ladder_expr_scope.
|
||||
Notation "'λ' x ',' τ '↦' e" := (expr_abs x τ e)
|
||||
(in custom ladder_expr at level 70, x constr, τ custom ladder_type at level 90, e custom ladder_expr at level 80, right associativity) :ladder_expr_scope.
|
||||
Notation "'λ' x ',' τ '↦morph' e" := (expr_morph x τ e)
|
||||
(in custom ladder_expr at level 70, x constr, τ custom ladder_type at level 90, e custom ladder_expr at level 80, right associativity) :ladder_expr_scope.
|
||||
Notation "'let' x ':=' e 'in' t" := (expr_let x e t)
|
||||
(in custom ladder_expr at level 20, x constr, e custom ladder_expr at level 99, t custom ladder_expr at level 99) : ladder_expr_scope.
|
||||
(in custom ladder_expr at level 60, x constr, e custom ladder_expr at level 80, t custom ladder_expr at level 80, right associativity) : ladder_expr_scope.
|
||||
Notation "e 'as' τ" := (expr_ascend τ e)
|
||||
(in custom ladder_expr at level 30, e custom ladder_expr, τ custom ladder_type at level 99) : ladder_expr_scope.
|
||||
Notation "e 'des' τ" := (expr_descend τ e)
|
||||
(in custom ladder_expr at level 30, e custom ladder_expr, τ custom ladder_type at level 99) : ladder_expr_scope.
|
||||
Notation "e1 e2" := (expr_app e1 e2)
|
||||
(in custom ladder_expr at level 50) : ladder_expr_scope.
|
||||
(in custom ladder_expr at level 90, e2 custom ladder_expr at next level) : ladder_expr_scope.
|
||||
Notation "e '#' τ" := (expr_ty_app e τ)
|
||||
(in custom ladder_expr at level 80, τ custom ladder_type at level 101, left associativity) : ladder_expr_scope.
|
||||
Notation "'(' e ')'" := e
|
||||
(in custom ladder_expr at level 0) : ladder_expr_scope.
|
||||
(in custom ladder_expr, e custom ladder_expr at next level, left associativity) : ladder_expr_scope.
|
||||
|
||||
|
||||
Print Grammar constr.
|
||||
|
||||
(* EXAMPLES *)
|
||||
|
||||
Open Scope ladder_type_scope.
|
||||
Open Scope ladder_expr_scope.
|
||||
|
||||
Check [< ∀"α", (< $"Seq"$ %"α"% > ~ < $"List"$ %"α"% >) >].
|
||||
Check [< ∀"α", [%"α"%] ~ <$"List"$ %"α"%> >].
|
||||
|
||||
Definition polymorphic_identity1 : expr_term := [{ Λ"T" ↦ λ"x"%"T"% ↦ (%"x"%) }].
|
||||
Definition polymorphic_identity2 : expr_term := [{ Λ"T" ↦ λ"y"%"T"% ↦ %"y"% }].
|
||||
|
||||
Definition polymorphic_identity1 : expr_term := [{ Λ"T" ↦ λ"x",%"T"% ↦ (%"x"%) }].
|
||||
Definition polymorphic_identity2 : expr_term := [{ Λ"T" ↦ λ"y",%"T"% ↦ %"y"% }].
|
||||
|
||||
Compute polymorphic_identity1.
|
||||
|
||||
|
|
110
coq/typing.v
110
coq/typing.v
|
@ -11,40 +11,41 @@ Require Import morph.
|
|||
|
||||
(** Typing Derivation *)
|
||||
|
||||
Open Scope ladder_expr_scope.
|
||||
|
||||
Reserved Notation "Gamma '|-' x '\is' X" (at level 101, x at next level, X at level 0).
|
||||
|
||||
Inductive expr_type : context -> expr_term -> type_term -> Prop :=
|
||||
| T_Var : forall Γ x τ,
|
||||
(context_contains Γ x τ) ->
|
||||
(Γ |- (expr_var x) \is τ)
|
||||
context_contains Γ x τ ->
|
||||
Γ |- [{ %x% }] \is τ
|
||||
|
||||
| T_Let : forall Γ s σ t τ x,
|
||||
(Γ |- s \is σ) ->
|
||||
((ctx_assign x σ Γ) |- t \is τ) ->
|
||||
(Γ |- (expr_let x s t) \is τ)
|
||||
Γ |- s \is σ ->
|
||||
(ctx_assign x σ Γ) |- t \is τ ->
|
||||
Γ |- [{ let x := s in t }] \is τ
|
||||
|
||||
| T_TypeAbs : forall Γ e τ α,
|
||||
Γ |- e \is τ ->
|
||||
Γ |- (expr_ty_abs α e) \is (type_univ α τ)
|
||||
Γ |- [{ Λ α ↦ e }] \is [< ∀α,τ >]
|
||||
|
||||
| T_TypeApp : forall Γ α e σ τ τ',
|
||||
Γ |- e \is (type_univ α τ) ->
|
||||
(type_subst1 α σ τ τ') ->
|
||||
Γ |- (expr_ty_app e σ) \is τ'
|
||||
| T_TypeApp : forall Γ α e σ τ,
|
||||
Γ |- e \is [< ∀α, τ >] ->
|
||||
Γ |- [{ e # σ }] \is (type_subst α σ τ)
|
||||
|
||||
| T_Abs : forall Γ x σ t τ,
|
||||
((ctx_assign x σ Γ) |- t \is τ) ->
|
||||
(Γ |- (expr_abs x σ t) \is (type_fun σ τ))
|
||||
(ctx_assign x σ Γ) |- t \is τ ->
|
||||
Γ |- [{ λ x , σ ↦ t }] \is [< σ -> τ >]
|
||||
|
||||
| T_MorphAbs : forall Γ x σ e τ,
|
||||
((ctx_assign x σ Γ) |- e \is τ) ->
|
||||
Γ |- (expr_morph x σ e) \is (type_morph σ τ)
|
||||
| T_MorphAbs : forall Γ x σ t τ,
|
||||
(ctx_assign x σ Γ) |- t \is τ ->
|
||||
Γ |- [{ λ x , σ ↦morph t }] \is [< σ ->morph τ >]
|
||||
|
||||
| T_App : forall Γ f a σ' σ τ,
|
||||
(Γ |- f \is (type_fun σ τ)) ->
|
||||
(Γ |- f \is [< σ -> τ >]) ->
|
||||
(Γ |- a \is σ') ->
|
||||
(Γ |- σ' ~> σ) ->
|
||||
(Γ |- (expr_app f a) \is τ)
|
||||
(Γ |- [{ (f a) }] \is τ)
|
||||
|
||||
| T_MorphFun : forall Γ f σ τ,
|
||||
Γ |- f \is (type_morph σ τ) ->
|
||||
|
@ -52,7 +53,7 @@ Inductive expr_type : context -> expr_term -> type_term -> Prop :=
|
|||
|
||||
| T_Ascend : forall Γ e τ τ',
|
||||
(Γ |- e \is τ) ->
|
||||
(Γ |- (expr_ascend τ' e) \is (type_ladder τ' τ))
|
||||
(Γ |- [{ e as τ' }] \is [< τ' ~ τ >])
|
||||
|
||||
| T_DescendImplicit : forall Γ x τ τ',
|
||||
Γ |- x \is τ ->
|
||||
|
@ -62,7 +63,7 @@ Inductive expr_type : context -> expr_term -> type_term -> Prop :=
|
|||
| T_Descend : forall Γ x τ τ',
|
||||
Γ |- x \is τ ->
|
||||
(τ :<= τ') ->
|
||||
Γ |- (expr_descend τ' x) \is τ'
|
||||
Γ |- [{ x des τ' }] \is τ'
|
||||
|
||||
where "Γ '|-' x '\is' τ" := (expr_type Γ x τ).
|
||||
|
||||
|
@ -75,73 +76,85 @@ Definition is_well_typed (e:expr_term) : Prop :=
|
|||
Inductive translate_typing : context -> expr_term -> type_term -> expr_term -> Prop :=
|
||||
| Expand_Var : forall Γ x τ,
|
||||
(Γ |- (expr_var x) \is τ) ->
|
||||
(translate_typing Γ (expr_var x) τ (expr_var x))
|
||||
(translate_typing Γ [{ %x% }] τ [{ %x% }])
|
||||
|
||||
| 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 Γ (expr_let x e t) τ (expr_let x e' t'))
|
||||
(translate_typing Γ
|
||||
[{ let x := e in t }]
|
||||
τ
|
||||
[{ let x := e' in t' }])
|
||||
|
||||
| Expand_TypeAbs : forall Γ α e e' τ,
|
||||
(Γ |- e \is τ) ->
|
||||
(translate_typing Γ e τ e') ->
|
||||
(translate_typing Γ
|
||||
(expr_ty_abs α e)
|
||||
(type_univ α τ)
|
||||
(expr_ty_abs α e'))
|
||||
[{ Λ α ↦ e }]
|
||||
[< ∀ α,τ >]
|
||||
[{ Λ α ↦ e' }])
|
||||
|
||||
| Expand_TypeApp : forall Γ e e' σ α τ,
|
||||
(Γ |- e \is (type_univ α τ)) ->
|
||||
(translate_typing Γ e τ e') ->
|
||||
(translate_typing Γ (expr_ty_app e τ) (type_subst α σ τ) (expr_ty_app e' τ))
|
||||
(translate_typing Γ
|
||||
[{ e # τ }]
|
||||
(type_subst α σ τ)
|
||||
[{ e' # τ }])
|
||||
|
||||
| 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 Γ (expr_abs x σ e) (type_fun σ τ) (expr_abs x σ e'))
|
||||
(translate_typing Γ
|
||||
[{ λ x , σ ↦ e }]
|
||||
[< σ -> τ >]
|
||||
[{ λ x , σ ↦ e' }])
|
||||
|
||||
| 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 Γ (expr_morph x σ e) (type_morph σ τ) (expr_morph x σ e'))
|
||||
(translate_typing Γ
|
||||
[{ λ x , σ ↦morph e }]
|
||||
[< σ ->morph τ >]
|
||||
[{ λ x , σ ↦morph e' }])
|
||||
|
||||
| Expand_App : forall Γ f f' a a' m σ τ σ',
|
||||
(Γ |- f \is (type_fun σ τ)) ->
|
||||
(Γ |- a \is σ') ->
|
||||
(Γ |- σ' ~> σ) ->
|
||||
(translate_typing Γ f (type_fun σ τ) f') ->
|
||||
(translate_typing Γ f [< σ -> τ >] f') ->
|
||||
(translate_typing Γ a σ' a') ->
|
||||
(translate_morphism_path Γ σ' σ m) ->
|
||||
(translate_typing Γ (expr_app f a) τ (expr_app f' (expr_app m a')))
|
||||
(translate_typing Γ [{ f a }] τ [{ f' (m a') }])
|
||||
|
||||
| Expand_MorphFun : forall Γ f f' σ τ,
|
||||
(Γ |- f \is (type_morph σ τ)) ->
|
||||
(Γ |- f \is (type_fun σ τ)) ->
|
||||
(translate_typing Γ f (type_morph σ τ) f') ->
|
||||
(translate_typing Γ f (type_fun σ τ) f')
|
||||
(translate_typing Γ f [< σ ->morph τ >] f') ->
|
||||
(translate_typing Γ f [< σ -> τ >] f')
|
||||
|
||||
| Expand_Ascend : forall Γ e e' τ τ',
|
||||
(Γ |- e \is τ) ->
|
||||
(Γ |- (expr_ascend τ' e) \is (type_ladder τ' τ)) ->
|
||||
(Γ |- [{ e as τ' }] \is [< τ' ~ τ >]) ->
|
||||
(translate_typing Γ e τ e') ->
|
||||
(translate_typing Γ (expr_ascend τ' e) (type_ladder τ' τ) (expr_ascend τ' e'))
|
||||
(translate_typing Γ [{ e as τ' }] [< τ' ~ τ >] [{ e' as τ' }])
|
||||
|
||||
| Expand_Descend : forall Γ e e' τ τ',
|
||||
(Γ |- e \is τ) ->
|
||||
(τ :<= τ') ->
|
||||
(Γ |- e \is τ') ->
|
||||
(Γ |- [{ e des τ' }] \is τ') ->
|
||||
(translate_typing Γ e τ e') ->
|
||||
(translate_typing Γ e τ' e')
|
||||
(translate_typing Γ e τ' [{ e' des τ' }])
|
||||
.
|
||||
|
||||
(* Examples *)
|
||||
|
||||
Example typing1 :
|
||||
ctx_empty |- [{ Λ"T" ↦ λ "x" %"T"% ↦ %"x"% }] \is [< ∀"T", %"T"% -> %"T"% >].
|
||||
ctx_empty |- [{ Λ"T" ↦ λ "x",%"T"% ↦ %"x"% }] \is [< ∀"T", %"T"% -> %"T"% >].
|
||||
Proof.
|
||||
intros.
|
||||
apply T_TypeAbs.
|
||||
|
@ -151,7 +164,7 @@ Proof.
|
|||
Qed.
|
||||
|
||||
Example typing2 :
|
||||
ctx_empty |- [{ Λ"T" ↦ λ "x" %"T"% ↦ %"x"% }] \is [< ∀"U", %"U"% -> %"U"% >].
|
||||
ctx_empty |- [{ Λ"T" ↦ λ "x",%"T"% ↦ %"x"% }] \is [< ∀"U", %"U"% -> %"U"% >].
|
||||
Proof.
|
||||
intros.
|
||||
apply T_DescendImplicit with (τ:=[< ∀"T",(%"T"% -> %"T"%) >]).
|
||||
|
@ -163,14 +176,11 @@ Proof.
|
|||
apply TSubRepr_Refl.
|
||||
apply TEq_Alpha.
|
||||
apply TAlpha_Rename.
|
||||
apply TSubst_Fun.
|
||||
apply TSubst_VarReplace.
|
||||
apply TSubst_VarReplace.
|
||||
Qed.
|
||||
|
||||
Example typing3 :
|
||||
ctx_empty |- [{
|
||||
Λ"T" ↦ Λ"U" ↦ λ"x" %"T"% ↦ λ"y" %"U"% ↦ %"y"%
|
||||
Λ"T" ↦ Λ"U" ↦ λ"x",%"T"% ↦ λ"y",%"U"% ↦ %"y"%
|
||||
}] \is [<
|
||||
∀"S",∀"T",(%"S"%->%"T"%->%"T"%)
|
||||
>].
|
||||
|
@ -186,26 +196,14 @@ Proof.
|
|||
apply TEq_Trans with (y:= [< ∀"S",∀"U",(%"S"%->%"U"%->%"U"%) >] ).
|
||||
apply TEq_Alpha.
|
||||
apply TAlpha_Rename.
|
||||
apply TSubst_UnivReplace. discriminate.
|
||||
easy.
|
||||
apply TSubst_Fun.
|
||||
apply TSubst_VarReplace.
|
||||
apply TSubst_Fun.
|
||||
apply TSubst_VarKeep. discriminate.
|
||||
apply TSubst_VarKeep. discriminate.
|
||||
|
||||
apply TEq_Alpha.
|
||||
apply TAlpha_SubUniv.
|
||||
apply TAlpha_Rename.
|
||||
apply TSubst_Fun.
|
||||
apply TSubst_VarKeep. discriminate.
|
||||
apply TSubst_Fun.
|
||||
apply TSubst_VarReplace.
|
||||
apply TSubst_VarReplace.
|
||||
Qed.
|
||||
|
||||
Example typing4 : (is_well_typed
|
||||
[{ Λ"T" ↦ Λ"U" ↦ λ"x" %"T"% ↦ λ"y" %"U"% ↦ %"x"% }]
|
||||
[{ Λ"T" ↦ Λ"U" ↦ λ"x",%"T"% ↦ λ"y",%"U"% ↦ %"x"% }]
|
||||
).
|
||||
Proof.
|
||||
unfold is_well_typed.
|
||||
|
@ -218,8 +216,6 @@ Proof.
|
|||
apply C_shuffle, C_take.
|
||||
Qed.
|
||||
|
||||
Open Scope ladder_expr_scope.
|
||||
|
||||
Example typing5 :
|
||||
(ctx_assign "60" [< $"ℝ"$ >]
|
||||
(ctx_assign "360" [< $"ℝ"$ >]
|
||||
|
@ -228,7 +224,7 @@ Example typing5 :
|
|||
|-
|
||||
[{
|
||||
let "deg2turns" :=
|
||||
(λ"x" $"Angle"$~$"Degrees"$~$"ℝ"$
|
||||
(λ"x",$"Angle"$~$"Degrees"$~$"ℝ"$
|
||||
↦morph (((%"/"% %"x"%) %"360"%) as $"Angle"$~$"Turns"$))
|
||||
in ( %"deg2turns"% (%"60"% as $"Angle"$~$"Degrees"$) )
|
||||
}]
|
||||
|
|
Loading…
Reference in a new issue