From Coq Require Import Lists.List. Require Import Atom. Require Import Environment. Require Import Metatheory. Require Import debruijn. Require Import subtype. Require Import env. Require Import morph. Require Import subst_lemmas. Require Import typing. Lemma transl_inv_let : forall Γ s s' σ t t' τ, (Γ |- s \is σ) -> (Γ |- [[ [{ let s in t }] \is τ ]] = [{ let s' in t' }]) -> forall L x, x `notin` L -> ((x, σ) :: Γ |- [[expr_open [{$ x}] t \is τ]] = expr_open [{$ x}] t') . Proof. Admitted. Lemma transl_inv_abs : forall Γ σ e e' τ, (Γ |- [[ [{ λ σ ↦ e }] \is τ ]] = [{ λ σ ↦ e' }]) -> forall L x, x `notin` L -> ((x, σ) :: Γ |- [[expr_open [{$ x}] e \is τ]] = expr_open [{$ x}] e') . Proof. Admitted. Lemma transl_inv_morph : forall Γ σ e e' τ, (Γ |- [[ [{ λ σ ↦morph e }] \is τ ]] = [{ λ σ ↦morph e' }]) -> forall L x, x `notin` L -> ((x, σ) :: Γ |- [[expr_open [{$ x}] e \is τ]] = expr_open [{$ x}] e') . Proof. Admitted. Lemma transl_inv_tabs : forall Γ e e' τ, (Γ |- [[ [{ Λ e }] \is [< ∀ τ >] ]] = [{ Λ e' }]) -> forall L x, x `notin` L -> (Γ |- [[ (expr_open_type (ty_fvar x) e) \is τ]] = expr_open_type (ty_fvar x) e') . Proof. Admitted.