From Coq Require Import Lists.List. Import ListNotations. 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 Γ s σ t τ x, (Γ |- s \is σ) -> (((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 Γ x σ t τ, (((x σ) :: Γ) |- t \is τ) -> (Γ |- [{ λ σ ↦ t }] \is [< σ -> τ >]) | T_MorphAbs : forall Γ x σ t τ, (((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 τ).