ladder-calculus/coq/typing_debruijn.v

142 lines
4.5 KiB
Coq
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

From Coq Require Import Lists.List.
Import ListNotations.
Require Import Metatheory.
Require Import Atom.
Require Import terms_debruijn.
Require Import subtype_debruijn.
Require Import context_debruijn.
Require Import morph_debruijn.
Open Scope ladder_type_scope.
Open Scope ladder_expr_scope.
Reserved Notation "Γ '|-' x '\is' X" (at level 101).
Inductive typing : context -> expr_DeBruijn -> type_DeBruijn -> Prop :=
| T_Var : forall Γ x τ,
(In (x, τ) Γ) ->
(Γ |- [{ $x }] \is τ)
| T_Let : forall Γ L s σ t τ,
(Γ |- s \is σ) ->
(forall x : atom, x `notin` L ->
( (x,σ)::Γ |- (expr_open (ex_fvar x) t) \is τ)) ->
(Γ |- [{ let s in t }] \is τ)
| T_TypeAbs : forall Γ e τ,
(Γ |- e \is τ) ->
(Γ |- [{ Λ e }] \is [< τ >])
| T_TypeApp : forall Γ e σ τ,
(Γ |- e \is [< τ >]) ->
(Γ |- [{ e # σ }] \is (type_open σ τ))
| T_Abs : forall Γ L σ t τ,
(forall x : atom, x `notin` L ->
( (x,σ)::Γ |- (expr_open (ex_fvar x) t) \is τ)) ->
(Γ |- [{ λ σ t }] \is [< σ -> τ >])
| T_MorphAbs : forall Γ L σ t τ,
(forall x : atom, x `notin` L ->
( (x,σ)::Γ |- (expr_open (ex_fvar x) t) \is τ)) ->
(Γ |- [{ λ σ morph t }] \is [< σ ->morph τ >])
| T_App : forall Γ f a σ' σ τ,
(Γ |- f \is [< σ -> τ >]) ->
(Γ |- a \is σ') ->
(Γ |- σ' ~~> σ) ->
(Γ |- [{ f a }] \is τ)
| T_MorphFun : forall Γ f σ τ,
(Γ |- f \is [< σ ->morph τ >]) ->
(Γ |- f \is [< σ -> τ >])
| T_Ascend : forall Γ e τ τ',
(Γ |- e \is τ) ->
(Γ |- [{ e as τ' }] \is [< τ' ~ τ >])
| T_DescendImplicit : forall Γ x τ τ',
(Γ |- x \is τ) ->
(τ :<= τ') ->
(Γ |- x \is τ')
| T_Descend : forall Γ x τ τ',
(Γ |- x \is τ) ->
(τ :<= τ') ->
(Γ |- [{ x des τ' }] \is τ')
where "Γ '|-' x '\is' τ" := (typing Γ x τ).
Reserved Notation "Γ '|-' '[[' e \is τ ']]' '=' f" (at level 101).
Inductive translate_typing : context -> expr_DeBruijn -> type_DeBruijn -> expr_DeBruijn -> Prop :=
| Expand_Var : forall Γ x τ,
(Γ |- [{ $x }] \is τ) ->
(Γ |- [[ [{ $x }] \is τ ]] = [{ $x }])
| Expand_Let : forall Γ L e e' t t' σ τ,
(Γ |- e \is σ) ->
(forall x : atom, x `notin` L ->
( (x,σ)::Γ |- (expr_open t (ex_fvar x)) \is τ) ->
( (x,σ)::Γ |- [[ t \is τ ]] = t' )
) ->
(Γ |- [[ e \is σ ]] = e') ->
(Γ |- [[ [{ let e in t }] \is τ ]] = [{ let e' in t' }])
| Expand_TypeAbs : forall Γ e e' τ,
(Γ |- e \is τ) ->
(Γ |- [[ e \is τ ]] = e') ->
(Γ |- [[ [{ Λ e }] \is [< τ >] ]] = [{ Λ e' }])
| Expand_TypeApp : forall Γ e e' σ τ,
(Γ |- e \is [< τ >]) ->
(Γ |- [[ e \is τ ]] = e') ->
(Γ |- [[ [{ e # σ }] \is (type_open σ τ) ]] = [{ e' # σ }])
| Expand_Abs : forall Γ L σ e e' τ,
(Γ |- [{ λ σ e }] \is [< σ -> τ >]) ->
(forall x : atom, x `notin` L ->
( (x,σ)::Γ |- (expr_open (ex_fvar x) e) \is τ) ->
( (x,σ)::Γ |- [[ e \is τ ]] = e' )
) ->
(Γ |- [[ [{ λ σ e }] \is [< σ -> τ >] ]] = [{ λ σ e' }])
| Expand_MorphAbs : forall Γ L σ e e' τ,
(Γ |- [{ λ σ e }] \is [< σ -> τ >]) ->
(forall x : atom, x `notin` L ->
( (x,σ)::Γ |- (expr_open (ex_fvar x) e) \is τ) ->
( (x,σ)::Γ |- [[ e \is τ ]] = e' )
) ->
(Γ |- [[ [{ λ σ morph e }] \is [< σ ->morph τ >] ]] = [{ λ σ morph e' }])
| Expand_App : forall Γ f f' a a' m σ τ σ',
(Γ |- f \is [< σ -> τ >]) ->
(Γ |- a \is σ') ->
(Γ |- σ' ~~> σ) ->
(Γ |- [[ f \is [< σ -> τ >] ]] = f') ->
(Γ |- [[ a \is σ' ]] = a') ->
(Γ |- [[ σ' ~~> σ ]] = m) ->
(Γ |- [[ [{ f a }] \is τ ]] = [{ f' (m a') }])
| Expand_MorphFun : forall Γ f f' σ τ,
(Γ |- f \is [< σ ->morph τ >]) ->
(Γ |- f \is [< σ -> τ >]) ->
(Γ |- [[ f \is [< σ ->morph τ >] ]] = f') ->
(Γ |- [[ f \is [< σ -> τ >] ]] = f')
| Expand_Ascend : forall Γ e e' τ τ',
(Γ |- e \is τ) ->
(Γ |- [{ e as τ' }] \is [< τ' ~ τ >]) ->
(Γ |- [[ e \is τ ]] = e') ->
(Γ |- [[ [{ e as τ' }] \is [< τ' ~ τ >] ]] = [{ e' as τ' }])
| Expand_Descend : forall Γ e e' τ τ',
(Γ |- e \is τ) ->
(τ :<= τ') ->
(Γ |- [{ e des τ' }] \is τ') ->
(Γ |- [[ e \is τ ]] = e') ->
(Γ |- [[ e \is τ' ]] = [{ e' des τ' }])
where "Γ '|-' '[[' e '\is' τ ']]' '=' f" := (translate_typing Γ e τ f).