Built with Alectryon, running Coq+SerAPI v8.16.0+0.16.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
From Tealeaves Require Export
Examples.SystemF.Syntax
Examples.SystemF.Contexts
Simplification.Tests.SystemF_LN.From Coq Require Import
Sorting.Permutation.Implicit Types (x : atom).Open Scope set_scope.(** * Properties of the typing judgment <<Δ ; Γ ⊢ t : τ>> *)(******************************************************************************)(** ** Structural properties <<Δ ; Γ ⊢ t : τ>> in <<Δ>> *)(******************************************************************************)(** *** <<Δ>> is always well-formed *)(******************************************************************************)
forall (Δ : kind_ctx) (Γ : type_ctx) (t : term LN)
(τ : typ LN), Δ; Γ ⊢ t : τ -> ok_kind_ctx Δ
forall (Δ : kind_ctx) (Γ : type_ctx) (t : term LN)
(τ : typ LN), Δ; Γ ⊢ t : τ -> ok_kind_ctx Δ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN e0: atom H: Δ ++ e0 ~ tt; Γ ⊢ open term ktyp (ty_v (Fr e0)) t
: open typ ktyp (ty_v (Fr e0)) τ e: atom H0: uniq (Δ ++ e ~ tt) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
uniq Δ
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 H0: Δ; Γ ⊢ t : ty_univ τ2 IHJudgment: uniq Δ
uniq Δ
all: (autorewrite with tea_rw_uniq in *; intuition).Qed.(** *** Permutation *)(******************************************************************************)
forall (Δ1Δ2 : list (atom * unit)) (Γ : type_ctx)
(t : term LN) (τ : typ LN),
Permutation Δ1 Δ2 -> Δ1; Γ ⊢ t : τ -> Δ2; Γ ⊢ t : τ
forall (Δ1Δ2 : list (atom * unit)) (Γ : type_ctx)
(t : term LN) (τ : typ LN),
Permutation Δ1 Δ2 -> Δ1; Γ ⊢ t : τ -> Δ2; Γ ⊢ t : τ
Δ1, Δ2: list (atom * unit) Γ: type_ctx t: term LN τ: typ LN perm: Permutation Δ1 Δ2 j: Δ1; Γ ⊢ t : τ
Δ2; Γ ⊢ t : τ
Δ1: list (atom * unit) Γ: type_ctx t: term LN τ: typ LN j: Δ1; Γ ⊢ t : τ
forallΔ2 : list (atom * unit),
Permutation Δ1 Δ2 -> Δ2; Γ ⊢ t : τ
Δ: kind_ctx Γ: type_ctx x: atom τ: typ LN H: ok_kind_ctx Δ H0: ok_type_ctx Δ Γ H1: (x, τ) ∈ Γ Δ2: list (atom * unit) Hperm: Permutation Δ Δ2
Δ2; Γ ⊢ tm_var (Fr x) : τ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 H0: forallx : AtomSet.elt,
x `notin` L ->
forallΔ2 : list (atom * unit),
Permutation Δ Δ2 ->
Δ2; Γ ++ x ~ τ1
⊢ open term ktrm (tm_var (Fr x)) t : τ2 Δ2: list (atom * unit) Hperm: Permutation Δ Δ2
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ H0: forallx : AtomSet.elt,
x `notin` L ->
forallΔ2 : list (atom * unit),
Permutation (Δ ++ x ~ tt) Δ2 ->
Δ2; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ Δ2: list (atom * unit) Hperm: Permutation Δ Δ2
Δ2; Γ ⊢ tm_tab t : ty_univ τ
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: forallΔ2 : list (atom * unit),
Permutation Δ Δ2 -> Δ2; Γ ⊢ t : ty_univ τ2 Δ2: list (atom * unit) Hperm: Permutation Δ Δ2
Δ2; Γ ⊢ tm_tap t τ1 : open typ ktyp τ1 τ2
Δ: kind_ctx Γ: type_ctx x: atom τ: typ LN H: ok_kind_ctx Δ H0: ok_type_ctx Δ Γ H1: (x, τ) ∈ Γ Δ2: list (atom * unit) Hperm: Permutation Δ Δ2
Δ2; Γ ⊢ tm_var (Fr x) : τ
eauto using Judgment, Permutation_sym with sysf_ctx.
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 H0: forallx : AtomSet.elt,
x `notin` L ->
forallΔ2 : list (atom * unit),
Permutation Δ Δ2 ->
Δ2; Γ ++ x ~ τ1
⊢ open term ktrm (tm_var (Fr x)) t : τ2 Δ2: list (atom * unit) Hperm: Permutation Δ Δ2
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ H0: forallx : AtomSet.elt,
x `notin` L ->
forallΔ2 : list (atom * unit),
Permutation (Δ ++ x ~ tt) Δ2 ->
Δ2; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ Δ2: list (atom * unit) Hperm: Permutation Δ Δ2
Δ2; Γ ⊢ tm_tab t : ty_univ τ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ H0: forallx : AtomSet.elt,
x `notin` L ->
forallΔ2 : list (atom * unit),
Permutation (Δ ++ x ~ tt) Δ2 ->
Δ2; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ Δ2: list (atom * unit) Hperm: Permutation Δ Δ2
forallx : AtomSet.elt,
x `notin` ?L ->
Δ2 ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ
eauto using Permutation_app_tail.
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: forallΔ2 : list (atom * unit),
Permutation Δ Δ2 -> Δ2; Γ ⊢ t : ty_univ τ2 Δ2: list (atom * unit) Hperm: Permutation Δ Δ2
Δ2; Γ ⊢ tm_tap t τ1 : open typ ktyp τ1 τ2
eauto using Judgment, ok_type_perm1.Qed.
forall (Δ1Δ2 : list (atom * unit)) x (Γ : type_ctx)
(t : term LN) (τ : typ LN),
(Δ1 ++ x ~ tt) ++ Δ2; Γ ⊢ t : τ ->
(Δ1 ++ Δ2) ++ x ~ tt; Γ ⊢ t : τ
forall (Δ1Δ2 : list (atom * unit)) x (Γ : type_ctx)
(t : term LN) (τ : typ LN),
(Δ1 ++ x ~ tt) ++ Δ2; Γ ⊢ t : τ ->
(Δ1 ++ Δ2) ++ x ~ tt; Γ ⊢ t : τ
Δ1, Δ2: list (atom * unit) x: atom Γ: type_ctx t: term LN τ: typ LN H: (Δ1 ++ x ~ tt) ++ Δ2; Γ ⊢ t : τ
(Δ1 ++ Δ2) ++ x ~ tt; Γ ⊢ t : τ
Δ1, Δ2: list (atom * unit) x: atom Γ: type_ctx t: term LN τ: typ LN H: (Δ1 ++ x ~ tt) ++ Δ2; Γ ⊢ t : τ
Permutation ?Δ1 ((Δ1 ++ Δ2) ++ x ~ tt)
Δ1, Δ2: list (atom * unit) x: atom Γ: type_ctx t: term LN τ: typ LN H: (Δ1 ++ x ~ tt) ++ Δ2; Γ ⊢ t : τ
?Δ1; Γ ⊢ t : τ
Δ1, Δ2: list (atom * unit) x: atom Γ: type_ctx t: term LN τ: typ LN H: (Δ1 ++ x ~ tt) ++ Δ2; Γ ⊢ t : τ
?Δ1; Γ ⊢ t : τ
eassumption.
Δ1, Δ2: list (atom * unit) x: atom Γ: type_ctx t: term LN τ: typ LN H: (Δ1 ++ x ~ tt) ++ Δ2; Γ ⊢ t : τ
Permutation ((Δ1 ++ x ~ tt) ++ Δ2)
((Δ1 ++ Δ2) ++ x ~ tt)
Δ1, Δ2: list (atom * unit) x: atom Γ: type_ctx t: term LN τ: typ LN H: (Δ1 ++ x ~ tt) ++ Δ2; Γ ⊢ t : τ
Permutation (Δ1 ++ x ~ tt ++ Δ2) (Δ1 ++ Δ2 ++ x ~ tt)
Δ1, Δ2: list (atom * unit) x: atom Γ: type_ctx t: term LN τ: typ LN H: (Δ1 ++ x ~ tt) ++ Δ2; Γ ⊢ t : τ
Permutation Δ1 Δ1
Δ1, Δ2: list (atom * unit) x: atom Γ: type_ctx t: term LN τ: typ LN H: (Δ1 ++ x ~ tt) ++ Δ2; Γ ⊢ t : τ
Permutation (x ~ tt ++ Δ2) (Δ2 ++ x ~ tt)
all: auto using Permutation_refl, Permutation_app_comm.Qed.(** *** Weakening *)(******************************************************************************)
Δ2: kind_ctx ok: ok_kind_ctx Δ2 Δ: kind_ctx disj: disjoint Δ Δ2 Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ H0: forallx : AtomSet.elt,
x `notin` L ->
disjoint (Δ ++ x ~ tt) Δ2 ->
(Δ ++ x ~ tt) ++ Δ2; Γ
⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ
Δ2: kind_ctx ok: ok_kind_ctx Δ2 Δ: kind_ctx disj: disjoint Δ Δ2 Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ H0: forallx : AtomSet.elt,
x `notin` L ->
disjoint (Δ ++ x ~ tt) Δ2 ->
(Δ ++ x ~ tt) ++ Δ2; Γ
⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ
Δ ++ Δ2; Γ ⊢ tm_tab t : ty_univ τ
Δ2: kind_ctx ok: ok_kind_ctx Δ2 Δ: kind_ctx disj: disjoint Δ Δ2 Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ IH: forallx : AtomSet.elt,
x `notin` L ->
disjoint (Δ ++ x ~ tt) Δ2 ->
(Δ ++ x ~ tt) ++ Δ2; Γ
⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ
Δ ++ Δ2; Γ ⊢ tm_tab t : ty_univ τ
Δ2: kind_ctx ok: ok_kind_ctx Δ2 Δ: kind_ctx disj: disjoint Δ Δ2 Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ IH: forallx : AtomSet.elt,
x `notin` L ->
disjoint (Δ ++ x ~ tt) Δ2 ->
(Δ ++ x ~ tt) ++ Δ2; Γ
⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ
forallx : AtomSet.elt,
x `notin` (L ∪ domset Δ2) ->
(Δ ++ Δ2) ++ x ~ tt; Γ
⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ
Δ2: kind_ctx ok: ok_kind_ctx Δ2 Δ: kind_ctx disj: disjoint Δ Δ2 Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: AtomSet.elt IH: disjoint (Δ ++ e ~ tt) Δ2 ->
(Δ ++ e ~ tt) ++ Δ2; Γ
⊢ open term ktyp (ty_v (Fr e)) t
: open typ ktyp (ty_v (Fr e)) τ Notin: e `notin` (L ∪ domset Δ2)
(Δ ++ Δ2) ++ e ~ tt; Γ
⊢ open term ktyp (ty_v (Fr e)) t
: open typ ktyp (ty_v (Fr e)) τ
Δ2: kind_ctx ok: ok_kind_ctx Δ2 Δ: kind_ctx disj: disjoint Δ Δ2 Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: AtomSet.elt IH: disjoint (Δ ++ e ~ tt) Δ2 ->
(Δ ++ e ~ tt) ++ Δ2; Γ
⊢ open term ktyp (ty_v (Fr e)) t
: open typ ktyp (ty_v (Fr e)) τ Notin: e `notin` (L ∪ domset Δ2)
(Δ ++ e ~ tt) ++ Δ2; Γ
⊢ open term ktyp (ty_v (Fr e)) t
: open typ ktyp (ty_v (Fr e)) τ
Δ2: kind_ctx ok: ok_kind_ctx Δ2 Δ: kind_ctx disj: disjoint Δ Δ2 Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: AtomSet.elt IH: disjoint Δ Δ2 /\ e `notin` domset Δ2 ->
(Δ ++ e ~ tt) ++ Δ2; Γ
⊢ open term ktyp (ty_v (Fr e)) t
: open typ ktyp (ty_v (Fr e)) τ Notin: e `notin` (L ∪ domset Δ2)
(Δ ++ e ~ tt) ++ Δ2; Γ
⊢ open term ktyp (ty_v (Fr e)) t
: open typ ktyp (ty_v (Fr e)) τ
apply j_inst; auto with sysf_ctx.Qed.(** *** Substitution *)(******************************************************************************)
forall (Δ1 : list (atom * unit)) x
(Δ2 : list (atom * unit)) (Γ : type_ctx)
(t : term LN) (ττ' : typ LN),
ok_type (Δ1 ++ Δ2) τ' ->
Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t : τ ->
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t : subst typ ktyp x τ' τ
forall (Δ1 : list (atom * unit)) x
(Δ2 : list (atom * unit)) (Γ : type_ctx)
(t : term LN) (ττ' : typ LN),
ok_type (Δ1 ++ Δ2) τ' ->
Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t : τ ->
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t : subst typ ktyp x τ' τ
Δ1: list (atom * unit) x: atom Δ2: list (atom * unit) Γ: type_ctx t: term LN τ, τ': typ LN Hok: ok_type (Δ1 ++ Δ2) τ' j: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t : τ
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t : subst typ ktyp x τ' τ
Δ1: list (atom * unit) x: atom Δ2: list (atom * unit) Γ: type_ctx t: term LN τ, τ': typ LN Hok: ok_type (Δ1 ++ Δ2) τ' rem: list (atom * unit) Heqrem: rem = Δ1 ++ x ~ tt ++ Δ2 j: rem; Γ ⊢ t : τ
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t : subst typ ktyp x τ' τ
Δ1: list (atom * unit) x: atom Γ: type_ctx t: term LN τ, τ': typ LN rem: list (atom * unit) j: rem; Γ ⊢ t : τ
forallΔ2 : list (atom * unit),
ok_type (Δ1 ++ Δ2) τ' ->
rem = Δ1 ++ x ~ tt ++ Δ2 ->
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t :
subst typ ktyp x τ' τ
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx x0: atom τ: typ LN Δ2: list (atom * unit) H0: ok_type_ctx (Δ1 ++ x ~ tt ++ Δ2) Γ H: ok_kind_ctx (Δ1 ++ x ~ tt ++ Δ2) H1: (x0, τ) ∈ Γ Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' (tm_var (Fr x0))
: subst typ ktyp x τ' τ
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN Δ2: list (atom * unit) H0: forallx0 : AtomSet.elt,
x0 `notin` L ->
forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3;
envmap (subst typ ktyp x τ') (Γ ++ x0 ~ τ1)
⊢ subst term ktyp x τ'
(open term ktrm (tm_var (Fr x0)) t)
: subst typ ktyp x τ' τ2 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ1 ++ x ~ tt ++ Δ2; Γ ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ2 Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' (tm_abs τ1 t)
: subst typ ktyp x τ' (ty_ar τ1 τ2)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx t1, t2: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IHj2: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t2
: subst typ ktyp x τ' τ1 IHj1: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t1
: subst typ ktyp x τ' (ty_ar τ1 τ2) j2: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t2 : τ1 j1: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t1 : ty_ar τ1 τ2 Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' (tm_app t1 t2)
: subst typ ktyp x τ' τ2
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN Δ2: list (atom * unit) H0: forallx0 : AtomSet.elt,
x0 `notin` L ->
forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt =
Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ'
(open term ktyp (ty_v (Fr x0)) t)
: subst typ ktyp x τ'
(open typ ktyp (ty_v (Fr x0)) τ) H: forallx0 : AtomSet.elt,
x0 `notin` L ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt; Γ
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' (tm_tab t)
: subst typ ktyp x τ' (ty_univ τ)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx t: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IHj: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t
: subst typ ktyp x τ' (ty_univ τ2) j: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t : ty_univ τ2 H: ok_type (Δ1 ++ x ~ tt ++ Δ2) τ1 Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' (tm_tap t τ1)
: subst typ ktyp x τ' (open typ ktyp τ1 τ2)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx x0: atom τ: typ LN Δ2: list (atom * unit) H0: ok_type_ctx (Δ1 ++ x ~ tt ++ Δ2) Γ H: ok_kind_ctx (Δ1 ++ x ~ tt ++ Δ2) H1: (x0, τ) ∈ Γ Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' (tm_var (Fr x0))
: subst typ ktyp x τ' τ
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx x0: atom τ: typ LN Δ2: list (atom * unit) H0: ok_type_ctx (Δ1 ++ x ~ tt ++ Δ2) Γ H: ok_kind_ctx (Δ1 ++ x ~ tt ++ Δ2) H1: (x0, τ) ∈ Γ Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ tm_var (Fr x0) : subst typ ktyp x τ' τ
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx x0: atom τ: typ LN Δ2: list (atom * unit) H0: ok_type_ctx (Δ1 ++ x ~ tt ++ Δ2) Γ H: ok_kind_ctx (Δ1 ++ x ~ tt ++ Δ2) H1: (x0, τ) ∈ Γ Hok: ok_type (Δ1 ++ Δ2) τ'
ok_kind_ctx (Δ1 ++ Δ2)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx x0: atom τ: typ LN Δ2: list (atom * unit) H0: ok_type_ctx (Δ1 ++ x ~ tt ++ Δ2) Γ H: ok_kind_ctx (Δ1 ++ x ~ tt ++ Δ2) H1: (x0, τ) ∈ Γ Hok: ok_type (Δ1 ++ Δ2) τ'
ok_type_ctx (Δ1 ++ Δ2)
(envmap (subst typ ktyp x τ') Γ)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx x0: atom τ: typ LN Δ2: list (atom * unit) H0: ok_type_ctx (Δ1 ++ x ~ tt ++ Δ2) Γ H: ok_kind_ctx (Δ1 ++ x ~ tt ++ Δ2) H1: (x0, τ) ∈ Γ Hok: ok_type (Δ1 ++ Δ2) τ'
(x0, subst typ ktyp x τ' τ)
∈ envmap (subst typ ktyp x τ') Γ
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx x0: atom τ: typ LN Δ2: list (atom * unit) H0: ok_type_ctx (Δ1 ++ x ~ tt ++ Δ2) Γ H: ok_kind_ctx (Δ1 ++ x ~ tt ++ Δ2) H1: (x0, τ) ∈ Γ Hok: ok_type (Δ1 ++ Δ2) τ'
ok_kind_ctx (Δ1 ++ Δ2)
eauto with sysf_ctx.
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx x0: atom τ: typ LN Δ2: list (atom * unit) H0: ok_type_ctx (Δ1 ++ x ~ tt ++ Δ2) Γ H: ok_kind_ctx (Δ1 ++ x ~ tt ++ Δ2) H1: (x0, τ) ∈ Γ Hok: ok_type (Δ1 ++ Δ2) τ'
ok_type_ctx (Δ1 ++ Δ2)
(envmap (subst typ ktyp x τ') Γ)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx x0: atom τ: typ LN Δ2: list (atom * unit) H0: ok_type_ctx (Δ1 ++ x ~ tt ++ Δ2) Γ H: ok_kind_ctx (Δ1 ++ x ~ tt ++ Δ2) H1: (x0, τ) ∈ Γ Hok: ok_type (Δ1 ++ Δ2) τ'
ok_type_ctx (Δ1 ++ x ~ tt ++ Δ2) Γ
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx x0: atom τ: typ LN Δ2: list (atom * unit) H0: ok_type_ctx (Δ1 ++ x ~ tt ++ Δ2) Γ H: ok_kind_ctx (Δ1 ++ x ~ tt ++ Δ2) H1: (x0, τ) ∈ Γ Hok: ok_type (Δ1 ++ Δ2) τ'
ok_type (Δ1 ++ Δ2) τ'
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx x0: atom τ: typ LN Δ2: list (atom * unit) H0: ok_type_ctx (Δ1 ++ x ~ tt ++ Δ2) Γ H: ok_kind_ctx (Δ1 ++ x ~ tt ++ Δ2) H1: (x0, τ) ∈ Γ Hok: ok_type (Δ1 ++ Δ2) τ'
LC typ ktyp τ'
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx x0: atom τ: typ LN Δ2: list (atom * unit) H0: ok_type_ctx (Δ1 ++ x ~ tt ++ Δ2) Γ H: ok_kind_ctx (Δ1 ++ x ~ tt ++ Δ2) H1: (x0, τ) ∈ Γ Hok: ok_type (Δ1 ++ Δ2) τ'
ok_type_ctx (Δ1 ++ x ~ tt ++ Δ2) Γ
eauto.
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx x0: atom τ: typ LN Δ2: list (atom * unit) H0: ok_type_ctx (Δ1 ++ x ~ tt ++ Δ2) Γ H: ok_kind_ctx (Δ1 ++ x ~ tt ++ Δ2) H1: (x0, τ) ∈ Γ Hok: ok_type (Δ1 ++ Δ2) τ'
ok_type (Δ1 ++ Δ2) τ'
eauto.
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx x0: atom τ: typ LN Δ2: list (atom * unit) H0: ok_type_ctx (Δ1 ++ x ~ tt ++ Δ2) Γ H: ok_kind_ctx (Δ1 ++ x ~ tt ++ Δ2) H1: (x0, τ) ∈ Γ Hok: ok_type (Δ1 ++ Δ2) τ'
LC typ ktyp τ'
eapply ok_type_lc; eauto.
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx x0: atom τ: typ LN Δ2: list (atom * unit) H0: ok_type_ctx (Δ1 ++ x ~ tt ++ Δ2) Γ H: ok_kind_ctx (Δ1 ++ x ~ tt ++ Δ2) H1: (x0, τ) ∈ Γ Hok: ok_type (Δ1 ++ Δ2) τ'
(x0, subst typ ktyp x τ' τ)
∈ envmap (subst typ ktyp x τ') Γ
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx x0: atom τ: typ LN Δ2: list (atom * unit) H0: ok_type_ctx (Δ1 ++ x ~ tt ++ Δ2) Γ H: ok_kind_ctx (Δ1 ++ x ~ tt ++ Δ2) H1: (x0, τ) ∈ Γ Hok: ok_type (Δ1 ++ Δ2) τ'
existsa : typ LN,
(x0, a) ∈ Γ /\
subst typ ktyp x τ' a = subst typ ktyp x τ' τ
eauto.
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN Δ2: list (atom * unit) H0: forallx0 : AtomSet.elt,
x0 `notin` L ->
forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3;
envmap (subst typ ktyp x τ') (Γ ++ x0 ~ τ1)
⊢ subst term ktyp x τ'
(open term ktrm (tm_var (Fr x0)) t)
: subst typ ktyp x τ' τ2 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ1 ++ x ~ tt ++ Δ2; Γ ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ2 Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' (tm_abs τ1 t)
: subst typ ktyp x τ' (ty_ar τ1 τ2)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IH: forallx0 : AtomSet.elt,
x0 `notin` L ->
forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3;
envmap (subst typ ktyp x τ') (Γ ++ x0 ~ τ1)
⊢ subst term ktyp x τ'
(open term ktrm (tm_var (Fr x0)) t)
: subst typ ktyp x τ' τ2 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ1 ++ x ~ tt ++ Δ2; Γ ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ2 Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' (tm_abs τ1 t)
: subst typ ktyp x τ' (ty_ar τ1 τ2)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IH: forallx0 : AtomSet.elt,
x0 `notin` L ->
forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3;
envmap (subst typ ktyp x τ') (Γ ++ x0 ~ τ1)
⊢ subst term ktyp x τ'
(open term ktrm (tm_var (Fr x0)) t)
: subst typ ktyp x τ' τ2 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ1 ++ x ~ tt ++ Δ2; Γ ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ2 Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ tm_abs (subst typ ktyp x τ' τ1)
(subst term ktyp x τ' t)
: ty_ar (subst typ ktyp x τ' τ1)
(subst typ ktyp x τ' τ2)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IH: forallx0 : AtomSet.elt,
x0 `notin` L ->
forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3;
envmap (subst typ ktyp x τ') (Γ ++ x0 ~ τ1)
⊢ subst term ktyp x τ'
(open term ktrm (tm_var (Fr x0)) t)
: subst typ ktyp x τ' τ2 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ1 ++ x ~ tt ++ Δ2; Γ ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ2 Hok: ok_type (Δ1 ++ Δ2) τ'
forallx0 : AtomSet.elt,
x0 `notin` (L ∪ {{x}}) ->
Δ1 ++ Δ2;
envmap (subst typ ktyp x τ') Γ ++
x0 ~ subst typ ktyp x τ' τ1
⊢ open term ktrm (tm_var (Fr x0))
(subst term ktyp x τ' t) :
subst typ ktyp x τ' τ2
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN Δ2: list (atom * unit) e: AtomSet.elt IH: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3;
envmap (subst typ ktyp x τ') (Γ ++ e ~ τ1)
⊢ subst term ktyp x τ'
(open term ktrm (tm_var (Fr e)) t)
: subst typ ktyp x τ' τ2 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ1 ++ x ~ tt ++ Δ2; Γ ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ2 Hok: ok_type (Δ1 ++ Δ2) τ' Notin: e `notin` (L ∪ {{x}})
Δ1 ++ Δ2;
envmap (subst typ ktyp x τ') Γ ++
e ~ subst typ ktyp x τ' τ1
⊢ open term ktrm (tm_var (Fr e))
(subst term ktyp x τ' t) :
subst typ ktyp x τ' τ2
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN Δ2: list (atom * unit) e: AtomSet.elt IH: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3;
envmap (subst typ ktyp x τ') (Γ ++ e ~ τ1)
⊢ open term ktrm
(subst (SystemF ktrm) ktyp x τ'
(tm_var (Fr e)))
(subst term ktyp x τ' t)
: subst typ ktyp x τ' τ2 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ1 ++ x ~ tt ++ Δ2; Γ ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ2 Hok: ok_type (Δ1 ++ Δ2) τ' Notin: e `notin` (L ∪ {{x}})
Δ1 ++ Δ2;
envmap (subst typ ktyp x τ') Γ ++
e ~ subst typ ktyp x τ' τ1
⊢ open term ktrm (tm_var (Fr e))
(subst term ktyp x τ' t) :
subst typ ktyp x τ' τ2
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN Δ2: list (atom * unit) e: AtomSet.elt IH: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3;
envmap (subst typ ktyp x τ') (Γ ++ e ~ τ1)
⊢ open term ktrm (tm_var (Fr e))
(subst term ktyp x τ' t)
: subst typ ktyp x τ' τ2 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ1 ++ x ~ tt ++ Δ2; Γ ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ2 Hok: ok_type (Δ1 ++ Δ2) τ' Notin: e `notin` (L ∪ {{x}})
Δ1 ++ Δ2;
envmap (subst typ ktyp x τ') Γ ++
e ~ subst typ ktyp x τ' τ1
⊢ open term ktrm (tm_var (Fr e))
(subst term ktyp x τ' t) :
subst typ ktyp x τ' τ2
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN Δ2: list (atom * unit) e: AtomSet.elt IH: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3;
envmap (subst typ ktyp x τ') Γ ++
envmap (subst typ ktyp x τ') (e ~ τ1)
⊢ open term ktrm (tm_var (Fr e))
(subst term ktyp x τ' t)
: subst typ ktyp x τ' τ2 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ1 ++ x ~ tt ++ Δ2; Γ ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ2 Hok: ok_type (Δ1 ++ Δ2) τ' Notin: e `notin` (L ∪ {{x}})
Δ1 ++ Δ2;
envmap (subst typ ktyp x τ') Γ ++
e ~ subst typ ktyp x τ' τ1
⊢ open term ktrm (tm_var (Fr e))
(subst term ktyp x τ' t) :
subst typ ktyp x τ' τ2
auto.
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx t1, t2: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IHj2: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t2
: subst typ ktyp x τ' τ1 IHj1: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t1
: subst typ ktyp x τ' (ty_ar τ1 τ2) j2: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t2 : τ1 j1: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t1 : ty_ar τ1 τ2 Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' (tm_app t1 t2)
: subst typ ktyp x τ' τ2
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx t1, t2: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IHj2: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t2
: subst typ ktyp x τ' τ1 IHj1: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t1
: subst typ ktyp x τ' (ty_ar τ1 τ2) j2: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t2 : τ1 j1: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t1 : ty_ar τ1 τ2 Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ tm_app (subst term ktyp x τ' t1)
(subst term ktyp x τ' t2) :
subst typ ktyp x τ' τ2
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx t1, t2: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IHj2: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t2
: subst typ ktyp x τ' τ1 IHj1: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t1
: subst typ ktyp x τ' (ty_ar τ1 τ2) j2: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t2 : τ1 j1: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t1 : ty_ar τ1 τ2 Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t1
: ty_ar (subst typ ktyp x τ' τ1)
(subst typ ktyp x τ' τ2)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx t1, t2: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IHj2: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t2
: subst typ ktyp x τ' τ1 IHj1: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t1
: subst typ ktyp x τ' (ty_ar τ1 τ2) j2: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t2 : τ1 j1: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t1 : ty_ar τ1 τ2 Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t2 :
subst typ ktyp x τ' τ1
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx t1, t2: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IHj2: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t2
: subst typ ktyp x τ' τ1 IHj1: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t1
: subst typ ktyp x τ' (ty_ar τ1 τ2) j2: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t2 : τ1 j1: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t1 : ty_ar τ1 τ2 Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t1
: ty_ar (subst typ ktyp x τ' τ1)
(subst typ ktyp x τ' τ2)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx t1, t2: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IHj2: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t2
: subst typ ktyp x τ' τ1 IHj1: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t1
: subst typ ktyp x τ' (ty_ar τ1 τ2) j2: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t2 : τ1 j1: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t1 : ty_ar τ1 τ2 Hok: ok_type (Δ1 ++ Δ2) τ'
ok_type (Δ1 ++ Δ2) τ'
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx t1, t2: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IHj2: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t2
: subst typ ktyp x τ' τ1 IHj1: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t1
: subst typ ktyp x τ' (ty_ar τ1 τ2) j2: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t2 : τ1 j1: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t1 : ty_ar τ1 τ2 Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ2
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx t1, t2: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IHj2: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t2
: subst typ ktyp x τ' τ1 IHj1: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t1
: subst typ ktyp x τ' (ty_ar τ1 τ2) j2: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t2 : τ1 j1: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t1 : ty_ar τ1 τ2 Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ2
reflexivity.
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx t1, t2: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IHj2: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t2
: subst typ ktyp x τ' τ1 IHj1: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t1
: subst typ ktyp x τ' (ty_ar τ1 τ2) j2: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t2 : τ1 j1: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t1 : ty_ar τ1 τ2 Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t2 :
subst typ ktyp x τ' τ1
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx t1, t2: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IHj2: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t2
: subst typ ktyp x τ' τ1 IHj1: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t1
: subst typ ktyp x τ' (ty_ar τ1 τ2) j2: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t2 : τ1 j1: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t1 : ty_ar τ1 τ2 Hok: ok_type (Δ1 ++ Δ2) τ'
ok_type (Δ1 ++ Δ2) τ'
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx t1, t2: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IHj2: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t2
: subst typ ktyp x τ' τ1 IHj1: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t1
: subst typ ktyp x τ' (ty_ar τ1 τ2) j2: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t2 : τ1 j1: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t1 : ty_ar τ1 τ2 Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ2
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx t1, t2: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IHj2: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t2
: subst typ ktyp x τ' τ1 IHj1: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t1
: subst typ ktyp x τ' (ty_ar τ1 τ2) j2: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t2 : τ1 j1: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t1 : ty_ar τ1 τ2 Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ2
reflexivity.
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN Δ2: list (atom * unit) H0: forallx0 : AtomSet.elt,
x0 `notin` L ->
forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt =
Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ'
(open term ktyp (ty_v (Fr x0)) t)
: subst typ ktyp x τ'
(open typ ktyp (ty_v (Fr x0)) τ) H: forallx0 : AtomSet.elt,
x0 `notin` L ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt; Γ
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' (tm_tab t)
: subst typ ktyp x τ' (ty_univ τ)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN Δ2: list (atom * unit) IH: forallx0 : AtomSet.elt,
x0 `notin` L ->
forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt =
Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ'
(open term ktyp (ty_v (Fr x0)) t)
: subst typ ktyp x τ'
(open typ ktyp (ty_v (Fr x0)) τ) H: forallx0 : AtomSet.elt,
x0 `notin` L ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt; Γ
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' (tm_tab t)
: subst typ ktyp x τ' (ty_univ τ)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN Δ2: list (atom * unit) IH: forallx0 : AtomSet.elt,
x0 `notin` L ->
forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt =
Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ'
(open term ktyp (ty_v (Fr x0)) t)
: subst typ ktyp x τ'
(open typ ktyp (ty_v (Fr x0)) τ) H: forallx0 : AtomSet.elt,
x0 `notin` L ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt; Γ
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ tm_tab (subst term ktyp x τ' t)
: ty_univ (subst typ ktyp x τ' τ)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN Δ2: list (atom * unit) IH: forallx0 : AtomSet.elt,
x0 `notin` L ->
forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt =
Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ'
(open term ktyp (ty_v (Fr x0)) t)
: subst typ ktyp x τ'
(open typ ktyp (ty_v (Fr x0)) τ) H: forallx0 : AtomSet.elt,
x0 `notin` L ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt; Γ
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ Hok: ok_type (Δ1 ++ Δ2) τ'
forallx0 : AtomSet.elt,
x0 `notin` (L ∪ {{x}}) ->
(Δ1 ++ Δ2) ++ x0 ~ tt; envmap (subst typ ktyp x τ') Γ
⊢ open term ktyp (ty_v (Fr x0))
(subst term ktyp x τ' t)
: open typ ktyp (ty_v (Fr x0)) (subst typ ktyp x τ' τ)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN Δ2: list (atom * unit) e: AtomSet.elt IH: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
(Δ1 ++ x ~ tt ++ Δ2) ++ e ~ tt =
Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ'
(open term ktyp (ty_v (Fr e)) t)
: subst typ ktyp x τ'
(open typ ktyp (ty_v (Fr e)) τ) H: forallx0 : AtomSet.elt,
x0 `notin` L ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt; Γ
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ Hok: ok_type (Δ1 ++ Δ2) τ' Notin: e `notin` (L ∪ {{x}})
(Δ1 ++ Δ2) ++ e ~ tt; envmap (subst typ ktyp x τ') Γ
⊢ open term ktyp (ty_v (Fr e))
(subst term ktyp x τ' t)
: open typ ktyp (ty_v (Fr e)) (subst typ ktyp x τ' τ)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN Δ2: list (atom * unit) e: AtomSet.elt IH: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 ++ e ~ tt = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ'
(open term ktyp (ty_v (Fr e)) t)
: subst typ ktyp x τ'
(open typ ktyp (ty_v (Fr e)) τ) H: forallx0 : AtomSet.elt,
x0 `notin` L ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt; Γ
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ Hok: ok_type (Δ1 ++ Δ2) τ' Notin: e `notin` (L ∪ {{x}})
Δ1 ++ Δ2 ++ e ~ tt; envmap (subst typ ktyp x τ') Γ
⊢ open term ktyp (ty_v (Fr e))
(subst term ktyp x τ' t)
: open typ ktyp (ty_v (Fr e)) (subst typ ktyp x τ' τ)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN Δ2: list (atom * unit) e: AtomSet.elt IH: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 ++ e ~ tt = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ'
(open term ktyp (ty_v (Fr e)) t)
: subst typ ktyp x τ'
(open typ ktyp (ty_v (Fr e)) τ) H: forallx0 : AtomSet.elt,
x0 `notin` L ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt; Γ
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ Hok: ok_type (Δ1 ++ Δ2) τ' Notin: e `notin` (L ∪ {{x}}) H0: x <> e
Δ1 ++ Δ2 ++ e ~ tt; envmap (subst typ ktyp x τ') Γ
⊢ open term ktyp (ty_v (Fr e))
(subst term ktyp x τ' t)
: open typ ktyp (ty_v (Fr e)) (subst typ ktyp x τ' τ)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN Δ2: list (atom * unit) e: AtomSet.elt IH: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 ++ e ~ tt = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ open term ktyp
(subst (SystemF ktyp) ktyp x τ' (ty_v (Fr e)))
(subst term ktyp x τ' t)
: subst typ ktyp x τ'
(open typ ktyp (ty_v (Fr e)) τ) H: forallx0 : AtomSet.elt,
x0 `notin` L ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt; Γ
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ Hok: ok_type (Δ1 ++ Δ2) τ' Notin: e `notin` (L ∪ {{x}}) H0: x <> e
Δ1 ++ Δ2 ++ e ~ tt; envmap (subst typ ktyp x τ') Γ
⊢ open term ktyp (ty_v (Fr e))
(subst term ktyp x τ' t)
: open typ ktyp (ty_v (Fr e)) (subst typ ktyp x τ' τ)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN Δ2: list (atom * unit) e: AtomSet.elt IH: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 ++ e ~ tt = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ'
(open term ktyp (ty_v (Fr e)) t)
: subst typ ktyp x τ'
(open typ ktyp (ty_v (Fr e)) τ) H: forallx0 : AtomSet.elt,
x0 `notin` L ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt; Γ
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ Hok: ok_type (Δ1 ++ Δ2) τ' Notin: e `notin` (L ∪ {{x}}) H0: x <> e
LC (SystemF ktyp) ktyp τ'
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN Δ2: list (atom * unit) e: AtomSet.elt IH: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 ++ e ~ tt = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ'
(open term ktyp (ty_v (Fr e)) t)
: subst typ ktyp x τ'
(open typ ktyp (ty_v (Fr e)) τ) H: forallx0 : AtomSet.elt,
x0 `notin` L ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt; Γ
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ Hok: ok_type (Δ1 ++ Δ2) τ' Notin: e `notin` (L ∪ {{x}}) H0: x <> e
LC (SystemF ktyp) ktyp τ'
now inverts Hok.
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN Δ2: list (atom * unit) e: AtomSet.elt IH: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 ++ e ~ tt = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ open term ktyp
(subst (SystemF ktyp) ktyp x τ' (ty_v (Fr e)))
(subst term ktyp x τ' t)
: subst typ ktyp x τ'
(open typ ktyp (ty_v (Fr e)) τ) H: forallx0 : AtomSet.elt,
x0 `notin` L ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt; Γ
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ Hok: ok_type (Δ1 ++ Δ2) τ' Notin: e `notin` (L ∪ {{x}}) H0: x <> e
Δ1 ++ Δ2 ++ e ~ tt; envmap (subst typ ktyp x τ') Γ
⊢ open term ktyp (ty_v (Fr e))
(subst term ktyp x τ' t)
: open typ ktyp (ty_v (Fr e)) (subst typ ktyp x τ' τ)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN Δ2: list (atom * unit) e: AtomSet.elt IH: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 ++ e ~ tt = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ open term ktyp
(subst (SystemF ktyp) ktyp x τ' (ty_v (Fr e)))
(subst term ktyp x τ' t)
: open typ ktyp
(subst (SystemF ktyp) ktyp x τ' (ty_v (Fr e)))
(subst typ ktyp x τ' τ) H: forallx0 : AtomSet.elt,
x0 `notin` L ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt; Γ
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ Hok: ok_type (Δ1 ++ Δ2) τ' Notin: e `notin` (L ∪ {{x}}) H0: x <> e
Δ1 ++ Δ2 ++ e ~ tt; envmap (subst typ ktyp x τ') Γ
⊢ open term ktyp (ty_v (Fr e))
(subst term ktyp x τ' t)
: open typ ktyp (ty_v (Fr e)) (subst typ ktyp x τ' τ)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN Δ2: list (atom * unit) e: AtomSet.elt IH: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 ++ e ~ tt = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ open term ktyp
(subst (SystemF ktyp) ktyp x τ' (ty_v (Fr e)))
(subst term ktyp x τ' t)
: subst typ ktyp x τ'
(open typ ktyp (ty_v (Fr e)) τ) H: forallx0 : AtomSet.elt,
x0 `notin` L ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt; Γ
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ Hok: ok_type (Δ1 ++ Δ2) τ' Notin: e `notin` (L ∪ {{x}}) H0: x <> e
LC (SystemF ktyp) ktyp τ'
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN Δ2: list (atom * unit) e: AtomSet.elt IH: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 ++ e ~ tt = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ open term ktyp
(subst (SystemF ktyp) ktyp x τ' (ty_v (Fr e)))
(subst term ktyp x τ' t)
: subst typ ktyp x τ'
(open typ ktyp (ty_v (Fr e)) τ) H: forallx0 : AtomSet.elt,
x0 `notin` L ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt; Γ
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ Hok: ok_type (Δ1 ++ Δ2) τ' Notin: e `notin` (L ∪ {{x}}) H0: x <> e
LC (SystemF ktyp) ktyp τ'
now inverts Hok.
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN Δ2: list (atom * unit) e: AtomSet.elt IH: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 ++ e ~ tt = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ open term ktyp
(subst (SystemF ktyp) ktyp x τ' (ty_v (Fr e)))
(subst term ktyp x τ' t)
: open typ ktyp
(subst (SystemF ktyp) ktyp x τ' (ty_v (Fr e)))
(subst typ ktyp x τ' τ) H: forallx0 : AtomSet.elt,
x0 `notin` L ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt; Γ
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ Hok: ok_type (Δ1 ++ Δ2) τ' Notin: e `notin` (L ∪ {{x}}) H0: x <> e
Δ1 ++ Δ2 ++ e ~ tt; envmap (subst typ ktyp x τ') Γ
⊢ open term ktyp (ty_v (Fr e))
(subst term ktyp x τ' t)
: open typ ktyp (ty_v (Fr e)) (subst typ ktyp x τ' τ)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN Δ2: list (atom * unit) e: AtomSet.elt IH: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 ++ e ~ tt = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ open term ktyp (ty_v (Fr e))
(subst term ktyp x τ' t)
: open typ ktyp (ty_v (Fr e))
(subst typ ktyp x τ' τ) H: forallx0 : AtomSet.elt,
x0 `notin` L ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt; Γ
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ Hok: ok_type (Δ1 ++ Δ2) τ' Notin: e `notin` (L ∪ {{x}}) H0: x <> e
Δ1 ++ Δ2 ++ e ~ tt; envmap (subst typ ktyp x τ') Γ
⊢ open term ktyp (ty_v (Fr e))
(subst term ktyp x τ' t)
: open typ ktyp (ty_v (Fr e)) (subst typ ktyp x τ' τ)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN Δ2: list (atom * unit) e: AtomSet.elt IH: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 ++ e ~ tt = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ open term ktyp (ty_v (Fr e))
(subst term ktyp x τ' t)
: open typ ktyp (ty_v (Fr e))
(subst typ ktyp x τ' τ) H: forallx0 : AtomSet.elt,
x0 `notin` L ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt; Γ
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ Hok: ok_type (Δ1 ++ Δ2) τ' Notin: e `notin` (L ∪ {{x}}) H0: x <> e
ok_type (Δ1 ++ Δ2 ++ e ~ tt) τ'
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN Δ2: list (atom * unit) e: AtomSet.elt IH: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 ++ e ~ tt = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ open term ktyp (ty_v (Fr e))
(subst term ktyp x τ' t)
: open typ ktyp (ty_v (Fr e))
(subst typ ktyp x τ' τ) H: forallx0 : AtomSet.elt,
x0 `notin` L ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt; Γ
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ Hok: ok_type (Δ1 ++ Δ2) τ' Notin: e `notin` (L ∪ {{x}}) H0: x <> e
Δ1 ++ x ~ tt ++ Δ2 ++ e ~ tt =
Δ1 ++ x ~ tt ++ Δ2 ++ e ~ tt
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN Δ2: list (atom * unit) e: AtomSet.elt IH: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 ++ e ~ tt = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ open term ktyp (ty_v (Fr e))
(subst term ktyp x τ' t)
: open typ ktyp (ty_v (Fr e))
(subst typ ktyp x τ' τ) H: forallx0 : AtomSet.elt,
x0 `notin` L ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt; Γ
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ Hok: ok_type (Δ1 ++ Δ2) τ' Notin: e `notin` (L ∪ {{x}}) H0: x <> e
ok_type (Δ1 ++ Δ2 ++ e ~ tt) τ'
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN Δ2: list (atom * unit) e: AtomSet.elt IH: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 ++ e ~ tt = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ open term ktyp (ty_v (Fr e))
(subst term ktyp x τ' t)
: open typ ktyp (ty_v (Fr e))
(subst typ ktyp x τ' τ) H: forallx0 : AtomSet.elt,
x0 `notin` L ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt; Γ
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ Hok: ok_type (Δ1 ++ Δ2) τ' Notin: e `notin` (L ∪ {{x}}) H0: x <> e
ok_type ((Δ1 ++ Δ2) ++ e ~ tt) τ'
auto using ok_type_weak_r.
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN Δ2: list (atom * unit) e: AtomSet.elt IH: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 ++ e ~ tt = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ open term ktyp (ty_v (Fr e))
(subst term ktyp x τ' t)
: open typ ktyp (ty_v (Fr e))
(subst typ ktyp x τ' τ) H: forallx0 : AtomSet.elt,
x0 `notin` L ->
(Δ1 ++ x ~ tt ++ Δ2) ++ x0 ~ tt; Γ
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ Hok: ok_type (Δ1 ++ Δ2) τ' Notin: e `notin` (L ∪ {{x}}) H0: x <> e
Δ1 ++ x ~ tt ++ Δ2 ++ e ~ tt =
Δ1 ++ x ~ tt ++ Δ2 ++ e ~ tt
reflexivity.
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx t: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IHj: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t
: subst typ ktyp x τ' (ty_univ τ2) j: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t : ty_univ τ2 H: ok_type (Δ1 ++ x ~ tt ++ Δ2) τ1 Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' (tm_tap t τ1)
: subst typ ktyp x τ' (open typ ktyp τ1 τ2)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx t: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IHj: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t
: subst typ ktyp x τ' (ty_univ τ2) j: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t : ty_univ τ2 H: ok_type (Δ1 ++ x ~ tt ++ Δ2) τ1 Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ tm_tap (subst term ktyp x τ' t)
(subst typ ktyp x τ' τ1)
: subst typ ktyp x τ' (open typ ktyp τ1 τ2)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx t: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IHj: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t
: ty_univ (subst typ ktyp x τ' τ2) j: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t : ty_univ τ2 H: ok_type (Δ1 ++ x ~ tt ++ Δ2) τ1 Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ tm_tap (subst term ktyp x τ' t)
(subst typ ktyp x τ' τ1)
: subst typ ktyp x τ' (open typ ktyp τ1 τ2)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx t: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IHj: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t
: subst typ ktyp x τ' (ty_univ τ2) j: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t : ty_univ τ2 H: ok_type (Δ1 ++ x ~ tt ++ Δ2) τ1 Hok: ok_type (Δ1 ++ Δ2) τ'
ty_univ (subst typ ktyp x τ' τ2) =
subst typ ktyp x τ' (ty_univ τ2)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx t: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IHj: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t
: subst typ ktyp x τ' (ty_univ τ2) j: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t : ty_univ τ2 H: ok_type (Δ1 ++ x ~ tt ++ Δ2) τ1 Hok: ok_type (Δ1 ++ Δ2) τ'
ty_univ (subst typ ktyp x τ' τ2) =
subst typ ktyp x τ' (ty_univ τ2)
now simplify_subst.
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx t: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IHj: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t
: ty_univ (subst typ ktyp x τ' τ2) j: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t : ty_univ τ2 H: ok_type (Δ1 ++ x ~ tt ++ Δ2) τ1 Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ tm_tap (subst term ktyp x τ' t)
(subst typ ktyp x τ' τ1)
: subst typ ktyp x τ' (open typ ktyp τ1 τ2)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx t: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IHj: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t
: ty_univ (subst typ ktyp x τ' τ2) j: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t : ty_univ τ2 H: ok_type (Δ1 ++ x ~ tt ++ Δ2) τ1 Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ tm_tap (subst term ktyp x τ' t)
(subst typ ktyp x τ' τ1)
: open typ ktyp (subst (SystemF ktyp) ktyp x τ' τ1)
(subst typ ktyp x τ' τ2)
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx t: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IHj: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t
: ty_univ (subst typ ktyp x τ' τ2) j: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t : ty_univ τ2 H: ok_type (Δ1 ++ x ~ tt ++ Δ2) τ1 Hok: ok_type (Δ1 ++ Δ2) τ'
LC (SystemF ktyp) ktyp τ'
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx t: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IHj: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t
: ty_univ (subst typ ktyp x τ' τ2) j: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t : ty_univ τ2 H: ok_type (Δ1 ++ x ~ tt ++ Δ2) τ1 Hok: ok_type (Δ1 ++ Δ2) τ'
LC (SystemF ktyp) ktyp τ'
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx t: term LN τ1, τ2: typ LN Δ2: list (atom * unit) H: ok_type (Δ1 ++ x ~ tt ++ Δ2) τ1 Hok: ok_type (Δ1 ++ Δ2) τ'
LC (SystemF ktyp) ktyp τ'
eapply ok_type_lc; eauto.
Δ1: list (atom * unit) x: atom τ': typ LN Γ: type_ctx t: term LN τ1, τ2: typ LN Δ2: list (atom * unit) IHj: forallΔ3 : list (atom * unit),
ok_type (Δ1 ++ Δ3) τ' ->
Δ1 ++ x ~ tt ++ Δ2 = Δ1 ++ x ~ tt ++ Δ3 ->
Δ1 ++ Δ3; envmap (subst typ ktyp x τ') Γ
⊢ subst term ktyp x τ' t
: ty_univ (subst typ ktyp x τ' τ2) j: Δ1 ++ x ~ tt ++ Δ2; Γ ⊢ t : ty_univ τ2 H: ok_type (Δ1 ++ x ~ tt ++ Δ2) τ1 Hok: ok_type (Δ1 ++ Δ2) τ'
Δ1 ++ Δ2; envmap (subst typ ktyp x τ') Γ
⊢ tm_tap (subst term ktyp x τ' t)
(subst typ ktyp x τ' τ1)
: open typ ktyp (subst (SystemF ktyp) ktyp x τ' τ1)
(subst typ ktyp x τ' τ2)
auto using j_inst with sysf_ctx.Qed.
forall (Δ : kind_ctx) x (Γ : type_ctx) (t : term LN)
(ττ2 : typ LN),
ok_type Δ τ2 ->
Δ ++ x ~ tt; Γ ⊢ t : τ ->
Δ; envmap (subst typ ktyp x τ2) Γ
⊢ subst term ktyp x τ2 t : subst typ ktyp x τ2 τ
forall (Δ : kind_ctx) x (Γ : type_ctx) (t : term LN)
(ττ2 : typ LN),
ok_type Δ τ2 ->
Δ ++ x ~ tt; Γ ⊢ t : τ ->
Δ; envmap (subst typ ktyp x τ2) Γ
⊢ subst term ktyp x τ2 t : subst typ ktyp x τ2 τ
Δ: kind_ctx x: atom Γ: type_ctx t: term LN τ, τ2: typ LN jτ: ok_type Δ τ2 jt: Δ ++ x ~ tt; Γ ⊢ t : τ
Δ; envmap (subst typ ktyp x τ2) Γ
⊢ subst term ktyp x τ2 t : subst typ ktyp x τ2 τ
Δ: kind_ctx x: atom Γ: type_ctx t: term LN τ, τ2: typ LN jτ: ok_type Δ τ2 jt: Δ ++ x ~ tt ++ []; Γ ⊢ t : τ
Δ; envmap (subst typ ktyp x τ2) Γ
⊢ subst term ktyp x τ2 t : subst typ ktyp x τ2 τ
Δ: kind_ctx x: atom Γ: type_ctx t: term LN τ, τ2: typ LN jτ: ok_type (Δ ++ []) τ2 jt: Δ ++ x ~ tt ++ []; Γ ⊢ t : τ
Δ; envmap (subst typ ktyp x τ2) Γ
⊢ subst term ktyp x τ2 t : subst typ ktyp x τ2 τ
Δ: kind_ctx x: atom Γ: type_ctx t: term LN τ, τ2: typ LN jτ: ok_type (Δ ++ []) τ2 jt: Δ ++ x ~ tt ++ []; Γ ⊢ t : τ
Δ ++ []; envmap (subst typ ktyp x τ2) Γ
⊢ subst term ktyp x τ2 t : subst typ ktyp x τ2 τ
auto using j_kind_ctx_subst.Qed.(** ** Properties of <<Δ ; Γ ⊢ t : τ>> in <<Γ>> *)(******************************************************************************)(** *** <<Γ>> is always well-formed *)(******************************************************************************)
forall (Δ : kind_ctx) (Γ : type_ctx) (t : term LN)
(τ : typ LN), Δ; Γ ⊢ t : τ -> ok_type_ctx Δ Γ
forall (Δ : kind_ctx) (Γ : type_ctx) (t : term LN)
(τ : typ LN), Δ; Γ ⊢ t : τ -> ok_type_ctx Δ Γ
Δ: kind_ctx Γ: type_ctx t: term LN τ: typ LN j: Δ; Γ ⊢ t : τ
ok_type_ctx Δ Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 H0: forallx : AtomSet.elt,
x `notin` L -> ok_type_ctx Δ (Γ ++ x ~ τ1)
ok_type_ctx Δ Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ H0: forallx : AtomSet.elt,
x `notin` L -> ok_type_ctx (Δ ++ x ~ tt) Γ
ok_type_ctx Δ Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 H0: forallx : AtomSet.elt,
x `notin` L -> ok_type_ctx Δ (Γ ++ x ~ τ1)
ok_type_ctx Δ Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom H0: ok_type_ctx Δ (Γ ++ e ~ τ1) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
ok_type_ctx Δ Γ
eauto using ok_type_ctx_inv_app_l.
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ H0: forallx : AtomSet.elt,
x `notin` L -> ok_type_ctx (Δ ++ x ~ tt) Γ
ok_type_ctx Δ Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ IH: forallx : AtomSet.elt,
x `notin` L -> ok_type_ctx (Δ ++ x ~ tt) Γ
ok_type_ctx Δ Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ IH: forallx : AtomSet.elt,
x `notin` L -> ok_type_ctx (Δ ++ x ~ tt) Γ e: atom Hfresh: e
`notin` (L
∪ atoms
(bind
(fun '(_, t) =>
free typ ktyp t) Γ))
ok_type_ctx Δ Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IH: ok_type_ctx (Δ ++ e ~ tt) Γ Hfresh: e
`notin` (L
∪ atoms
(bind
(fun '(_, t) =>
free typ ktyp t) Γ))
ok_type_ctx Δ Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IH: ok_type_ctx (Δ ++ e ~ tt) Γ Hfresh: e
`notin` (L
∪ atoms
(bind
(fun '(_, t) =>
free typ ktyp t) Γ))
ok_type_ctx (Δ ++ e ~ tt) Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IH: ok_type_ctx (Δ ++ e ~ tt) Γ Hfresh: e
`notin` (L
∪ atoms
(bind
(fun '(_, t) =>
free typ ktyp t) Γ))
forallt : typ LN,
t ∈ range Γ -> e `notin` FV typ ktyp t
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IH: ok_type_ctx (Δ ++ e ~ tt) Γ Hfresh: e
`notin` (L
∪ atoms
(bind
(fun '(_, t) =>
free typ ktyp t) Γ))
ok_type_ctx (Δ ++ e ~ tt) Γ
assumption.
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IH: ok_type_ctx (Δ ++ e ~ tt) Γ Hfresh: e
`notin` (L
∪ atoms
(bind
(fun '(_, t) =>
free typ ktyp t) Γ))
forallt : typ LN,
t ∈ range Γ -> e `notin` FV typ ktyp t
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IH: ok_type_ctx (Δ ++ e ~ tt) Γ Hfresh: e
`notin` (L
∪ atoms
(bind
(fun '(_, t) =>
free typ ktyp t) Γ)) t0: typ LN Hin: t0 ∈ range Γ
e `notin` FV typ ktyp t0
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IH: ok_type_ctx (Δ ++ e ~ tt) Γ Hfresh: e
`notin` (L
∪ atoms
(bind
(fun '(_, t) =>
free typ ktyp t) Γ)) t0: typ LN Hin: t0 ∈ range Γ lemma: ~ e ∈ bind (fun '(_, t) => free typ ktyp t) Γ
e `notin` FV typ ktyp t0
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IH: ok_type_ctx (Δ ++ e ~ tt) Γ Hfresh: e
`notin` (L
∪ atoms
(bind
(fun '(_, t) =>
free typ ktyp t) Γ)) t0: typ LN Hin: t0 ∈ range Γ
~ e ∈ bind (fun '(_, t) => free typ ktyp t) Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IH: ok_type_ctx (Δ ++ e ~ tt) Γ Hfresh: e
`notin` (L
∪ atoms
(bind
(fun '(_, t) =>
free typ ktyp t) Γ)) t0: typ LN Hin: t0 ∈ range Γ lemma: ~ e ∈ bind (fun '(_, t) => free typ ktyp t) Γ
e `notin` FV typ ktyp t0
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IH: ok_type_ctx (Δ ++ e ~ tt) Γ Hfresh: e
`notin` (L
∪ atoms
(bind
(fun '(_, t) =>
free typ ktyp t) Γ)) t0: typ LN Hin: t0 ∈ range Γ lemma: ~
(existsa : atom * typ LN,
a ∈ Γ /\
e ∈ (let '(_, t) := a in free typ ktyp t))
e `notin` FV typ ktyp t0
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IH: ok_type_ctx (Δ ++ e ~ tt) Γ Hfresh: e
`notin` (L
∪ atoms
(bind
(fun '(_, t) =>
free typ ktyp t) Γ)) t0: typ LN Hin: existsx, (x, t0) ∈ Γ lemma: ~
(existsa : atom * typ LN,
a ∈ Γ /\
e ∈ (let '(_, t) := a in free typ ktyp t))
e `notin` FV typ ktyp t0
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IH: ok_type_ctx (Δ ++ e ~ tt) Γ Hfresh: e
`notin` (L
∪ atoms
(bind
(fun '(_, t) =>
free typ ktyp t) Γ)) t0: typ LN x: atom Hin: (x, t0) ∈ Γ lemma: ~
(existsa : atom * typ LN,
a ∈ Γ /\
e ∈ (let '(_, t) := a in free typ ktyp t))
e `notin` FV typ ktyp t0
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IH: ok_type_ctx (Δ ++ e ~ tt) Γ Hfresh: e
`notin` (L
∪ atoms
(bind
(fun '(_, t) =>
free typ ktyp t) Γ)) t0: typ LN x: atom Hin: (x, t0) ∈ Γ lemma: e `in` FV typ ktyp t0
existsa : atom * typ LN,
a ∈ Γ /\ e ∈ (let '(_, t) := a in free typ ktyp t)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IH: ok_type_ctx (Δ ++ e ~ tt) Γ Hfresh: e
`notin` (L
∪ atoms
(bind
(fun '(_, t) =>
free typ ktyp t) Γ)) t0: typ LN x: atom Hin: (x, t0) ∈ Γ lemma: e `in` FV typ ktyp t0
(x, t0) ∈ Γ /\ e ∈ free typ ktyp t0
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IH: ok_type_ctx (Δ ++ e ~ tt) Γ Hfresh: e
`notin` (L
∪ atoms
(bind
(fun '(_, t) =>
free typ ktyp t) Γ)) t0: typ LN x: atom Hin: (x, t0) ∈ Γ lemma: e `in` FV typ ktyp t0
e ∈ free typ ktyp t0
nowapply free_iff_FV.
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IH: ok_type_ctx (Δ ++ e ~ tt) Γ Hfresh: e
`notin` (L
∪ atoms
(bind
(fun '(_, t) =>
free typ ktyp t) Γ)) t0: typ LN Hin: t0 ∈ range Γ
~ e ∈ bind (fun '(_, t) => free typ ktyp t) Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IH: ok_type_ctx (Δ ++ e ~ tt) Γ Hfresh: e
`notin` (L
∪ atoms
(bind
(fun '(_, t) =>
free typ ktyp t) Γ)) t0: typ LN Hin: t0 ∈ range Γ
~ e ∈ bind (fun '(_, t) => free typ ktyp t) Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IH: ok_type_ctx (Δ ++ e ~ tt) Γ Hfresh: e
`notin` (L
∪ atoms
(bind
(fun '(_, t) =>
free typ ktyp t) Γ)) t0: typ LN Hin: t0 ∈ range Γ
e
`notin` atoms
(bind (fun '(_, t) => free typ ktyp t) Γ)
forall (Δ : kind_ctx) (Γ1 : type_ctx)
(Γ2 : list (atom * typ LN)) (t : term LN)
(τ : typ LN),
Δ; Γ1 ⊢ t : τ -> Permutation Γ1 Γ2 -> Δ; Γ2 ⊢ t : τ
forall (Δ : kind_ctx) (Γ1 : type_ctx)
(Γ2 : list (atom * typ LN)) (t : term LN)
(τ : typ LN),
Δ; Γ1 ⊢ t : τ -> Permutation Γ1 Γ2 -> Δ; Γ2 ⊢ t : τ
Δ: kind_ctx Γ1: type_ctx Γ2: list (atom * typ LN) t: term LN τ: typ LN j: Δ; Γ1 ⊢ t : τ perm: Permutation Γ1 Γ2
Δ; Γ2 ⊢ t : τ
Δ: kind_ctx Γ1: type_ctx t: term LN τ: typ LN j: Δ; Γ1 ⊢ t : τ
forallΓ2 : list (atom * typ LN),
Permutation Γ1 Γ2 -> Δ; Γ2 ⊢ t : τ
Δ: kind_ctx Γ: type_ctx x: atom τ: typ LN H: ok_kind_ctx Δ H0: ok_type_ctx Δ Γ H1: (x, τ) ∈ Γ Γ2: list (atom * typ LN) Hperm: Permutation Γ Γ2
Δ; Γ2 ⊢ tm_var (Fr x) : τ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 H0: forallx : AtomSet.elt,
x `notin` L ->
forallΓ2 : list (atom * typ LN),
Permutation (Γ ++ x ~ τ1) Γ2 ->
Δ; Γ2 ⊢ open term ktrm (tm_var (Fr x)) t : τ2 Γ2: list (atom * typ LN) Hperm: Permutation Γ Γ2
Δ; Γ2 ⊢ tm_abs τ1 t : ty_ar τ1 τ2
Δ: kind_ctx Γ: type_ctx x: atom τ: typ LN H: ok_kind_ctx Δ H0: ok_type_ctx Δ Γ H1: (x, τ) ∈ Γ Γ2: list (atom * typ LN) Hperm: Permutation Γ Γ2
Δ; Γ2 ⊢ tm_var (Fr x) : τ
Δ: kind_ctx Γ: type_ctx x: atom τ: typ LN H: ok_kind_ctx Δ H0: ok_type_ctx Δ Γ H1: (x, τ) ∈ Γ Γ2: list (atom * typ LN) Hperm: Permutation Γ Γ2
ok_kind_ctx Δ
Δ: kind_ctx Γ: type_ctx x: atom τ: typ LN H: ok_kind_ctx Δ H0: ok_type_ctx Δ Γ H1: (x, τ) ∈ Γ Γ2: list (atom * typ LN) Hperm: Permutation Γ Γ2
ok_type_ctx Δ Γ2
Δ: kind_ctx Γ: type_ctx x: atom τ: typ LN H: ok_kind_ctx Δ H0: ok_type_ctx Δ Γ H1: (x, τ) ∈ Γ Γ2: list (atom * typ LN) Hperm: Permutation Γ Γ2
(x, τ) ∈ Γ2
Δ: kind_ctx Γ: type_ctx x: atom τ: typ LN H: ok_kind_ctx Δ H0: ok_type_ctx Δ Γ H1: (x, τ) ∈ Γ Γ2: list (atom * typ LN) Hperm: Permutation Γ Γ2
ok_kind_ctx Δ
eauto using ok_kind_ctx_perm.
Δ: kind_ctx Γ: type_ctx x: atom τ: typ LN H: ok_kind_ctx Δ H0: ok_type_ctx Δ Γ H1: (x, τ) ∈ Γ Γ2: list (atom * typ LN) Hperm: Permutation Γ Γ2
ok_type_ctx Δ Γ2
rewrite ok_type_ctx_perm_gamma;
eauto using Permutation_sym.
Δ: kind_ctx Γ: type_ctx x: atom τ: typ LN H: ok_kind_ctx Δ H0: ok_type_ctx Δ Γ H1: (x, τ) ∈ Γ Γ2: list (atom * typ LN) Hperm: Permutation Γ Γ2
(x, τ) ∈ Γ2
Δ: kind_ctx Γ: type_ctx x: atom τ: typ LN H: ok_kind_ctx Δ H0: ok_type_ctx Δ Γ H1: (x, τ) ∈ Γ Γ2: list (atom * typ LN) Hperm: Permutation Γ2 Γ
(x, τ) ∈ Γ2
erewrite List.permutation_spec; eauto.
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 H0: forallx : AtomSet.elt,
x `notin` L ->
forallΓ2 : list (atom * typ LN),
Permutation (Γ ++ x ~ τ1) Γ2 ->
Δ; Γ2 ⊢ open term ktrm (tm_var (Fr x)) t : τ2 Γ2: list (atom * typ LN) Hperm: Permutation Γ Γ2
Δ; Γ2 ⊢ tm_abs τ1 t : ty_ar τ1 τ2
econstructor; eauto using Permutation_app_tail.Qed.
forall (Δ : kind_ctx) (Γ1Γ2 : list (atom * typ LN))
(t : term LN) (τ : typ LN),
Permutation Γ1 Γ2 -> Δ; Γ1 ⊢ t : τ <-> Δ; Γ2 ⊢ t : τ
forall (Δ : kind_ctx) (Γ1Γ2 : list (atom * typ LN))
(t : term LN) (τ : typ LN),
Permutation Γ1 Γ2 -> Δ; Γ1 ⊢ t : τ <-> Δ; Γ2 ⊢ t : τ
Δ: kind_ctx Γ1, Γ2: list (atom * typ LN) t: term LN τ: typ LN Hperm: Permutation Γ1 Γ2
Δ; Γ1 ⊢ t : τ <-> Δ; Γ2 ⊢ t : τ
Δ: kind_ctx Γ1, Γ2: list (atom * typ LN) t: term LN τ: typ LN Hperm: Permutation Γ1 Γ2 H: Permutation Γ2 Γ1
Δ; Γ1 ⊢ t : τ <-> Δ; Γ2 ⊢ t : τ
split; eauto using j_type_ctx_perm.Qed.
forall (Δ : kind_ctx) (Γ1Γ2 : list (atom * typ LN))
x (τ2 : typ LN) (t : term LN) (τ1 : typ LN),
Δ; (Γ1 ++ x ~ τ2) ++ Γ2 ⊢ t : τ1 <->
Δ; (Γ1 ++ Γ2) ++ x ~ τ2 ⊢ t : τ1
forall (Δ : kind_ctx) (Γ1Γ2 : list (atom * typ LN))
x (τ2 : typ LN) (t : term LN) (τ1 : typ LN),
Δ; (Γ1 ++ x ~ τ2) ++ Γ2 ⊢ t : τ1 <->
Δ; (Γ1 ++ Γ2) ++ x ~ τ2 ⊢ t : τ1
Δ: kind_ctx Γ1, Γ2: list (atom * typ LN) x: atom τ2: typ LN t: term LN τ1: typ LN
Δ; (Γ1 ++ x ~ τ2) ++ Γ2 ⊢ t : τ1 <->
Δ; (Γ1 ++ Γ2) ++ x ~ τ2 ⊢ t : τ1
Δ: kind_ctx Γ1, Γ2: list (atom * typ LN) x: atom τ2: typ LN t: term LN τ1: typ LN
Permutation ((Γ1 ++ x ~ τ2) ++ Γ2)
((Γ1 ++ Γ2) ++ x ~ τ2)
Δ: kind_ctx Γ1, Γ2: list (atom * typ LN) x: atom τ2: typ LN t: term LN τ1: typ LN
Permutation (Γ1 ++ x ~ τ2 ++ Γ2) (Γ1 ++ Γ2 ++ x ~ τ2)
apply Permutation_app; auto using Permutation_app_comm.Qed.(** *** Weakening *)(******************************************************************************)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 H0: forallx : AtomSet.elt,
x `notin` L -> LC typ ktyp τ2
LC typ ktyp (ty_ar τ1 τ2)
Δ: kind_ctx Γ: type_ctx t1, t2: term LN τ1, τ2: typ LN j1: Δ; Γ ⊢ t1 : ty_ar τ1 τ2 j2: Δ; Γ ⊢ t2 : τ1 IHj1: LC typ ktyp (ty_ar τ1 τ2) IHj2: LC typ ktyp τ1
LC typ ktyp τ2
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ H0: forallx : AtomSet.elt,
x `notin` L ->
LC typ ktyp (open typ ktyp (ty_v (Fr x)) τ)
LC typ ktyp (ty_univ τ)
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: LC typ ktyp (ty_univ τ2)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 H0: forallx : AtomSet.elt,
x `notin` L -> LC typ ktyp τ2
LC typ ktyp (ty_ar τ1 τ2)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 H0: LC typ ktyp τ2 e: atom Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC typ ktyp (ty_ar τ1 τ2)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 H0: LC typ ktyp τ2 e: atom Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC typ ktyp τ1 /\ LC typ ktyp τ2
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 H0: LC typ ktyp τ2 e: atom Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC typ ktyp τ1
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 H0: LC typ ktyp τ2 e: atom Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC typ ktyp τ2
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 H0: LC typ ktyp τ2 e: atom Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC typ ktyp τ1
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: ok_type Δ τ1 H0: LC typ ktyp τ2 e: atom Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC typ ktyp τ1
eapply ok_type_lc; eauto.
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 H0: LC typ ktyp τ2 e: atom Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC typ ktyp τ2
auto.
Δ: kind_ctx Γ: type_ctx t1, t2: term LN τ1, τ2: typ LN j1: Δ; Γ ⊢ t1 : ty_ar τ1 τ2 j2: Δ; Γ ⊢ t2 : τ1 IHj1: LC typ ktyp (ty_ar τ1 τ2) IHj2: LC typ ktyp τ1
LC typ ktyp τ2
nowreplace (LC (ix := I2) typ ktyp (ty_ar τ1 τ2))
with (LC typ ktyp τ1 /\ LC typ ktyp τ2) in IHj1 by (now simplify_LC).
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ H0: forallx : AtomSet.elt,
x `notin` L ->
LC typ ktyp (open typ ktyp (ty_v (Fr x)) τ)
LC typ ktyp (ty_univ τ)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN e0: atom H: Δ ++ e0 ~ tt; Γ ⊢ open term ktyp (ty_v (Fr e0)) t
: open typ ktyp (ty_v (Fr e0)) τ e: atom H0: LC typ ktyp (open typ ktyp (ty_v (Fr e)) τ) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC typ ktyp (ty_univ τ)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN e0: atom H: Δ ++ e0 ~ tt; Γ ⊢ open term ktyp (ty_v (Fr e0)) t
: open typ ktyp (ty_v (Fr e0)) τ e: atom H0: LC typ ktyp (open typ ktyp (ty_v (Fr e)) τ) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LCn typ ktyp 1 τ
rewrite (open_lc_gap_eq_var typ); eauto.
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: LC typ ktyp (ty_univ τ2)
LC typ ktyp (open typ ktyp τ1 τ2)
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: LC typ ktyp (ty_univ τ2)
LCn typ ktyp 1 τ2
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: LC typ ktyp (ty_univ τ2)
LC (SystemF ktyp) ktyp τ1
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: LC typ ktyp (ty_univ τ2)
LCn typ ktyp 1 τ2
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: LCn typ ktyp 1 τ2
LCn typ ktyp 1 τ2
assumption.
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: LC typ ktyp (ty_univ τ2)
LC (SystemF ktyp) ktyp τ1
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: LC typ ktyp (ty_univ τ2)
ok_type ?Δ τ1
eassumption.Qed.(** *** <<τ>> is always well-formed *)(******************************************************************************)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 H0: forallx : AtomSet.elt,
x `notin` L -> ok_type Δ τ2
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ H0: forallx : AtomSet.elt,
x `notin` L ->
ok_type (Δ ++ x ~ tt)
(open typ ktyp (ty_v (Fr x)) τ)
ok_type Δ (ty_univ τ)
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: ok_type Δ (ty_univ τ2)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 H0: forallx : AtomSet.elt,
x `notin` L -> ok_type Δ τ2
ok_type Δ (ty_ar τ1 τ2)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 H0: ok_type Δ τ2 e: atom Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
ok_type Δ (ty_ar τ1 τ2)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 H0: ok_type Δ τ2 e: atom Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
ok_type Δ τ1 /\ ok_type Δ τ2
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 H0: ok_type Δ τ2 e: atom Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
ok_type Δ τ1
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 H0: ok_type Δ τ2 e: atom Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
ok_type Δ τ2
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 H0: ok_type Δ τ2 e: atom Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
ok_type Δ τ1
eapply j_type_ctx2; eauto.
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 H0: ok_type Δ τ2 e: atom Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ H0: forallx : AtomSet.elt,
x `notin` L ->
ok_type (Δ ++ x ~ tt)
(open typ ktyp (ty_v (Fr x)) τ)
ok_type Δ (ty_univ τ)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ H0: forallx : AtomSet.elt,
x `notin` L ->
ok_type (Δ ++ x ~ tt)
(open typ ktyp (ty_v (Fr x)) τ)
scoped typ ktyp (ty_univ τ) (domset Δ) /\
LCn typ ktyp 1 τ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ IH: forallx : AtomSet.elt,
x `notin` L ->
ok_type (Δ ++ x ~ tt)
(open typ ktyp (ty_v (Fr x)) τ)
scoped typ ktyp (ty_univ τ) (domset Δ) /\
LCn typ ktyp 1 τ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ IH: forallx : AtomSet.elt,
x `notin` L ->
ok_type (Δ ++ x ~ tt)
(open typ ktyp (ty_v (Fr x)) τ) e: atom Hfresh: e `notin` (L ∪ FV typ ktyp τ)
scoped typ ktyp (ty_univ τ) (domset Δ) /\
LCn typ ktyp 1 τ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IH: ok_type (Δ ++ e ~ tt)
(open typ ktyp (ty_v (Fr e)) τ) Hfresh: e `notin` (L ∪ FV typ ktyp τ)
scoped typ ktyp (ty_univ τ) (domset Δ) /\
LCn typ ktyp 1 τ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IHsc: scoped typ ktyp (open typ ktyp (ty_v (Fr e)) τ)
(domset (Δ ++ e ~ tt)) IHlc: LC typ ktyp (open typ ktyp (ty_v (Fr e)) τ) Hfresh: e `notin` (L ∪ FV typ ktyp τ)
scoped typ ktyp (ty_univ τ) (domset Δ) /\
LCn typ ktyp 1 τ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IHsc: scoped typ ktyp (open typ ktyp (ty_v (Fr e)) τ)
(domset (Δ ++ e ~ tt)) IHlc: LC typ ktyp (open typ ktyp (ty_v (Fr e)) τ) Hfresh: e `notin` (L ∪ FV typ ktyp τ)
scoped typ ktyp (ty_univ τ) (domset Δ)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IHsc: scoped typ ktyp (open typ ktyp (ty_v (Fr e)) τ)
(domset (Δ ++ e ~ tt)) IHlc: LC typ ktyp (open typ ktyp (ty_v (Fr e)) τ) Hfresh: e `notin` (L ∪ FV typ ktyp τ)
LCn typ ktyp 1 τ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IHsc: scoped typ ktyp (open typ ktyp (ty_v (Fr e)) τ)
(domset (Δ ++ e ~ tt)) IHlc: LC typ ktyp (open typ ktyp (ty_v (Fr e)) τ) Hfresh: e `notin` (L ∪ FV typ ktyp τ)
scoped typ ktyp (ty_univ τ) (domset Δ)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IHsc: FV typ ktyp (open typ ktyp (ty_v (Fr e)) τ)
⊆ domset (Δ ++ e ~ tt) IHlc: LC typ ktyp (open typ ktyp (ty_v (Fr e)) τ) Hfresh: e `notin` (L ∪ FV typ ktyp τ)
FV typ ktyp (ty_univ τ) ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IHsc: FV typ ktyp (open typ ktyp (ty_v (Fr e)) τ)
⊆ domset (Δ ++ e ~ tt) IHlc: LC typ ktyp (open typ ktyp (ty_v (Fr e)) τ) Hfresh: e `notin` (L ∪ FV typ ktyp τ) cut: FV typ ktyp τ ⊆ domset (Δ ++ e ~ tt)
FV typ ktyp (ty_univ τ) ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IHsc: FV typ ktyp (open typ ktyp (ty_v (Fr e)) τ)
⊆ domset (Δ ++ e ~ tt) IHlc: LC typ ktyp (open typ ktyp (ty_v (Fr e)) τ) Hfresh: e `notin` (L ∪ FV typ ktyp τ)
FV typ ktyp τ ⊆ domset (Δ ++ e ~ tt)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IHsc: FV typ ktyp (open typ ktyp (ty_v (Fr e)) τ)
⊆ domset (Δ ++ e ~ tt) IHlc: LC typ ktyp (open typ ktyp (ty_v (Fr e)) τ) Hfresh: e `notin` (L ∪ FV typ ktyp τ) cut: FV typ ktyp τ ⊆ domset (Δ ++ e ~ tt)
FV typ ktyp (ty_univ τ) ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IHsc: FV typ ktyp (open typ ktyp (ty_v (Fr e)) τ)
⊆ domset (Δ ++ e ~ tt) IHlc: LC typ ktyp (open typ ktyp (ty_v (Fr e)) τ) Hfresh: e `notin` (L ∪ FV typ ktyp τ) cut: FV typ ktyp τ ⊆ domset Δ ∪ {{e}}
FV typ ktyp (ty_univ τ) ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IHsc: FV typ ktyp (open typ ktyp (ty_v (Fr e)) τ)
⊆ domset (Δ ++ e ~ tt) IHlc: LC typ ktyp (open typ ktyp (ty_v (Fr e)) τ) Hfresh: e `notin` (L ∪ FV typ ktyp τ) cut: FV typ ktyp τ ⊆ domset Δ ∪ {{e}}
FV typ ktyp τ ⊆ domset Δ
fsetdec.
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IHsc: FV typ ktyp (open typ ktyp (ty_v (Fr e)) τ)
⊆ domset (Δ ++ e ~ tt) IHlc: LC typ ktyp (open typ ktyp (ty_v (Fr e)) τ) Hfresh: e `notin` (L ∪ FV typ ktyp τ)
FV typ ktyp τ ⊆ domset (Δ ++ e ~ tt)
rewrite FV_open_lower; eauto; typeclasses eauto.
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IHsc: scoped typ ktyp (open typ ktyp (ty_v (Fr e)) τ)
(domset (Δ ++ e ~ tt)) IHlc: LC typ ktyp (open typ ktyp (ty_v (Fr e)) τ) Hfresh: e `notin` (L ∪ FV typ ktyp τ)
LCn typ ktyp 1 τ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ e: atom IHsc: scoped typ ktyp (open typ ktyp (ty_v (Fr e)) τ)
(domset (Δ ++ e ~ tt)) IHlc: LC typ ktyp (open typ ktyp (ty_v (Fr e)) τ) Hfresh: e `notin` (L ∪ FV typ ktyp τ)
LC (SystemF ktyp) ktyp (ty_v (Fr e))
now simplify_LC.
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: ok_type Δ (ty_univ τ2)
ok_type Δ (open typ ktyp τ1 τ2)
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: scoped typ ktyp τ1 (domset Δ) /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: scoped typ ktyp (ty_univ τ2) (domset Δ) /\
LC typ ktyp (ty_univ τ2)
scoped typ ktyp (open typ ktyp τ1 τ2) (domset Δ) /\
LC typ ktyp (open typ ktyp τ1 τ2)
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: FV typ ktyp τ1 ⊆ domset Δ /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: FV typ ktyp (ty_univ τ2) ⊆ domset Δ /\
LC typ ktyp (ty_univ τ2)
FV typ ktyp (open typ ktyp τ1 τ2) ⊆ domset Δ /\
LC typ ktyp (open typ ktyp τ1 τ2)
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: FV typ ktyp τ1 ⊆ domset Δ /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj_FV: FV typ ktyp (ty_univ τ2) ⊆ domset Δ IHj_LC: LC typ ktyp (ty_univ τ2)
FV typ ktyp (open typ ktyp τ1 τ2) ⊆ domset Δ /\
LC typ ktyp (open typ ktyp τ1 τ2)
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: FV typ ktyp τ1 ⊆ domset Δ /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj_FV: FV typ ktyp (ty_univ τ2) ⊆ domset Δ IHj_LC: LC typ ktyp (ty_univ τ2)
FV typ ktyp (open typ ktyp τ1 τ2) ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: FV typ ktyp τ1 ⊆ domset Δ /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj_FV: FV typ ktyp (ty_univ τ2) ⊆ domset Δ IHj_LC: LC typ ktyp (ty_univ τ2)
LC typ ktyp (open typ ktyp τ1 τ2)
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: FV typ ktyp τ1 ⊆ domset Δ /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj_FV: FV typ ktyp (ty_univ τ2) ⊆ domset Δ IHj_LC: LC typ ktyp (ty_univ τ2)
FV typ ktyp (open typ ktyp τ1 τ2) ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: FV typ ktyp τ1 ⊆ domset Δ /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj_FV: FV typ ktyp (ty_univ τ2) ⊆ domset Δ IHj_LC: LC typ ktyp (ty_univ τ2)
FV typ ktyp (open typ ktyp τ1 τ2)
⊆ FV typ ktyp τ1 ∪ FV typ ktyp τ2
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: FV typ ktyp τ1 ⊆ domset Δ /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj_FV: FV typ ktyp (ty_univ τ2) ⊆ domset Δ IHj_LC: LC typ ktyp (ty_univ τ2)
FV typ ktyp τ1 ∪ FV typ ktyp τ2 ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: FV typ ktyp τ1 ⊆ domset Δ /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj_FV: FV typ ktyp (ty_univ τ2) ⊆ domset Δ IHj_LC: LC typ ktyp (ty_univ τ2)
FV typ ktyp (open typ ktyp τ1 τ2)
⊆ FV typ ktyp τ1 ∪ FV typ ktyp τ2
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: FV typ ktyp τ1 ⊆ domset Δ /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj_FV: FV typ ktyp (ty_univ τ2) ⊆ domset Δ IHj_LC: LC typ ktyp (ty_univ τ2)
FV typ ktyp τ2 ∪ FV (SystemF ktyp) ktyp τ1
⊆ FV typ ktyp τ1 ∪ FV typ ktyp τ2
fsetdec.
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: FV typ ktyp τ1 ⊆ domset Δ /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj_FV: FV typ ktyp (ty_univ τ2) ⊆ domset Δ IHj_LC: LC typ ktyp (ty_univ τ2)
FV typ ktyp τ1 ∪ FV typ ktyp τ2 ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: FV typ ktyp τ1 ⊆ domset Δ /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj_FV: FV typ ktyp τ2 ⊆ domset Δ IHj_LC: LC typ ktyp (ty_univ τ2)
FV typ ktyp τ1 ∪ FV typ ktyp τ2 ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: FV typ ktyp τ1 ⊆ domset Δ /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj_FV: FV typ ktyp (ty_univ τ2) ⊆ domset Δ IHj_LC: LC typ ktyp (ty_univ τ2)
FV typ ktyp τ2 = FV typ ktyp (ty_univ τ2)
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: FV typ ktyp τ1 ⊆ domset Δ /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj_FV: FV typ ktyp (ty_univ τ2) ⊆ domset Δ IHj_LC: LC typ ktyp (ty_univ τ2)
FV typ ktyp τ2 = FV typ ktyp (ty_univ τ2)
now simplify_FV.
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: FV typ ktyp τ1 ⊆ domset Δ /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj_FV: FV typ ktyp (ty_univ τ2) ⊆ domset Δ IHj_LC: LC typ ktyp (ty_univ τ2)
LC typ ktyp (open typ ktyp τ1 τ2)
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: FV typ ktyp τ1 ⊆ domset Δ /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj_FV: FV typ ktyp (ty_univ τ2) ⊆ domset Δ IHj_LC: LCn typ ktyp 1 τ2
LC typ ktyp (open typ ktyp τ1 τ2)
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: FV typ ktyp τ1 ⊆ domset Δ /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj_FV: FV typ ktyp (ty_univ τ2) ⊆ domset Δ IHj_LC: LC typ ktyp (ty_univ τ2)
LCn typ ktyp 1 τ2 = LC typ ktyp (ty_univ τ2)
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: FV typ ktyp τ1 ⊆ domset Δ /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj_FV: FV typ ktyp (ty_univ τ2) ⊆ domset Δ IHj_LC: LCn typ ktyp 1 τ2
LCn typ ktyp 1 τ2
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: FV typ ktyp τ1 ⊆ domset Δ /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj_FV: FV typ ktyp (ty_univ τ2) ⊆ domset Δ IHj_LC: LCn typ ktyp 1 τ2
LC (SystemF ktyp) ktyp τ1
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: FV typ ktyp τ1 ⊆ domset Δ /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj_FV: FV typ ktyp (ty_univ τ2) ⊆ domset Δ IHj_LC: LC typ ktyp (ty_univ τ2)
LCn typ ktyp 1 τ2 = LC typ ktyp (ty_univ τ2)
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: FV typ ktyp τ1 ⊆ domset Δ /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj_FV: FV typ ktyp (ty_univ τ2) ⊆ domset Δ IHj_LC: LCn typ ktyp 1 τ2
LCn typ ktyp 1 τ2
assumption.
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: FV typ ktyp τ1 ⊆ domset Δ /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj_FV: FV typ ktyp (ty_univ τ2) ⊆ domset Δ IHj_LC: LCn typ ktyp 1 τ2
LC (SystemF ktyp) ktyp τ1
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: FV typ ktyp τ1 ⊆ domset Δ /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj_FV: FV typ ktyp (ty_univ τ2) ⊆ domset Δ IHj_LC: LC typ ktyp (ty_univ τ2)
LCn typ ktyp 1 τ2 = LC typ ktyp (ty_univ τ2)
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: FV typ ktyp τ1 ⊆ domset Δ /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj_FV: FV typ ktyp (ty_univ τ2) ⊆ domset Δ IHj_LC: LCn typ ktyp 1 τ2
LC (SystemF ktyp) ktyp τ1
tauto.
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: FV typ ktyp τ1 ⊆ domset Δ /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj_FV: FV typ ktyp (ty_univ τ2) ⊆ domset Δ IHj_LC: LC typ ktyp (ty_univ τ2)
LCn typ ktyp 1 τ2 = LC typ ktyp (ty_univ τ2)
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: FV typ ktyp τ1 ⊆ domset Δ /\ LC typ ktyp τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj_FV: FV typ ktyp (ty_univ τ2) ⊆ domset Δ IHj_LC: LC typ ktyp (ty_univ τ2)
LCn typ ktyp 1 τ2 = LC typ ktyp (ty_univ τ2)
now simplify_LC.}Qed.(** ** Properties in t of <<(Δ ; Γ ⊢ t : τ>> *)(******************************************************************************)(** *** <<t>> is always locally closed w.r.t. term variables *)(******************************************************************************)
forall (Δ : kind_ctx) (Γ : type_ctx) (t : term LN)
(τ : typ LN), Δ; Γ ⊢ t : τ -> LC term ktrm t
forall (Δ : kind_ctx) (Γ : type_ctx) (t : term LN)
(τ : typ LN), Δ; Γ ⊢ t : τ -> LC term ktrm t
Δ: kind_ctx Γ: type_ctx t: term LN τ: typ LN j: Δ; Γ ⊢ t : τ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 IH: forallx : AtomSet.elt,
x `notin` L ->
LC term ktrm (open term ktrm (tm_var (Fr x)) t)
LC term ktrm (tm_abs τ1 t)
Δ: kind_ctx Γ: type_ctx t1, t2: term LN τ1, τ2: typ LN j1: Δ; Γ ⊢ t1 : ty_ar τ1 τ2 j2: Δ; Γ ⊢ t2 : τ1 IHj1: LC term ktrm t1 IHj2: LC term ktrm t2
LC term ktrm (tm_app t1 t2)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ IH: forallx : AtomSet.elt,
x `notin` L ->
LC term ktrm (open term ktyp (ty_v (Fr x)) t)
LC term ktrm (tm_tab t)
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: LC term ktrm t
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 IH: forallx : AtomSet.elt,
x `notin` L ->
LC term ktrm (open term ktrm (tm_var (Fr x)) t)
LC term ktrm (tm_abs τ1 t)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: LC term ktrm (open term ktrm (tm_var (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC term ktrm (tm_abs τ1 t)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: LC term ktrm (open term ktrm (tm_var (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC typ ktrm τ1 /\ LCn term ktrm 1 t
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: LC term ktrm (open term ktrm (tm_var (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC typ ktrm τ1
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: LC term ktrm (open term ktrm (tm_var (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LCn term ktrm 1 t
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: LC term ktrm (open term ktrm (tm_var (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC typ ktrm τ1
apply LC_typ_trm.
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: LC term ktrm (open term ktrm (tm_var (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LCn term ktrm 1 t
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: LC term ktrm (open term ktrm (tm_var (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC term ktrm (open term ktrm ?Goal3 t)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: LC term ktrm (open term ktrm (tm_var (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC (SystemF ktrm) ktrm ?Goal3
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: LC term ktrm (open term ktrm (tm_var (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC term ktrm (open term ktrm ?Goal3 t)
eassumption.
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: LC term ktrm (open term ktrm (tm_var (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC (SystemF ktrm) ktrm (tm_var (Fr e))
now simplify_LC.
Δ: kind_ctx Γ: type_ctx t1, t2: term LN τ1, τ2: typ LN j1: Δ; Γ ⊢ t1 : ty_ar τ1 τ2 j2: Δ; Γ ⊢ t2 : τ1 IHj1: LC term ktrm t1 IHj2: LC term ktrm t2
LC term ktrm (tm_app t1 t2)
now simplify_LC.
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ IH: forallx : AtomSet.elt,
x `notin` L ->
LC term ktrm (open term ktyp (ty_v (Fr x)) t)
LC term ktrm (tm_tab t)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN e0: atom H: Δ ++ e0 ~ tt; Γ ⊢ open term ktyp (ty_v (Fr e0)) t
: open typ ktyp (ty_v (Fr e0)) τ e: atom IH: LC term ktrm (open term ktyp (ty_v (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC term ktrm (tm_tab t)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN e0: atom H: Δ ++ e0 ~ tt; Γ ⊢ open term ktyp (ty_v (Fr e0)) t
: open typ ktyp (ty_v (Fr e0)) τ e: atom IH: LC term ktrm (open term ktyp (ty_v (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC term ktrm t
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN e0: atom H: Δ ++ e0 ~ tt; Γ ⊢ open term ktyp (ty_v (Fr e0)) t
: open typ ktyp (ty_v (Fr e0)) τ e: atom IH: LC term ktrm (open term ktyp (ty_v (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC term ktrm
(open term ?Goal1 (mret SystemF ?Goal1 (Fr ?Goal2))
t)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN e0: atom H: Δ ++ e0 ~ tt; Γ ⊢ open term ktyp (ty_v (Fr e0)) t
: open typ ktyp (ty_v (Fr e0)) τ e: atom IH: LC term ktrm (open term ktyp (ty_v (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
ktrm <> ?Goal1
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN e0: atom H: Δ ++ e0 ~ tt; Γ ⊢ open term ktyp (ty_v (Fr e0)) t
: open typ ktyp (ty_v (Fr e0)) τ e: atom IH: LC term ktrm (open term ktyp (ty_v (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC term ktrm
(open term ?Goal1 (mret SystemF ?Goal1 (Fr ?Goal2))
t)
exact IH.
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN e0: atom H: Δ ++ e0 ~ tt; Γ ⊢ open term ktyp (ty_v (Fr e0)) t
: open typ ktyp (ty_v (Fr e0)) τ e: atom IH: LC term ktrm (open term ktyp (ty_v (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
ktrm <> ktyp
discriminate.
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: LC term ktrm t
LC term ktrm (tm_tap t τ1)
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: LC term ktrm t
LC term ktrm t /\ LC typ ktrm τ1
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: LC term ktrm t
LC term ktrm t
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: LC term ktrm t
LC typ ktrm τ1
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: LC term ktrm t
LC term ktrm t
assumption.
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: LC term ktrm t
LC typ ktrm τ1
apply LC_typ_trm.Qed.(** *** <<t>> is always locally closed w.r.t. type variables *)(******************************************************************************)
forall (Δ : kind_ctx) (Γ : type_ctx) (t : term LN)
(τ : typ LN), Δ; Γ ⊢ t : τ -> LC term ktyp t
forall (Δ : kind_ctx) (Γ : type_ctx) (t : term LN)
(τ : typ LN), Δ; Γ ⊢ t : τ -> LC term ktyp t
Δ: kind_ctx Γ: type_ctx t: term LN τ: typ LN j: Δ; Γ ⊢ t : τ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 H0: forallx : AtomSet.elt,
x `notin` L ->
LC term ktyp (open term ktrm (tm_var (Fr x)) t)
LC term ktyp (tm_abs τ1 t)
Δ: kind_ctx Γ: type_ctx t1, t2: term LN τ1, τ2: typ LN j1: Δ; Γ ⊢ t1 : ty_ar τ1 τ2 j2: Δ; Γ ⊢ t2 : τ1 IHj1: LC term ktyp t1 IHj2: LC term ktyp t2
LC term ktyp (tm_app t1 t2)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ H0: forallx : AtomSet.elt,
x `notin` L ->
LC term ktyp (open term ktyp (ty_v (Fr x)) t)
LC term ktyp (tm_tab t)
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: LC term ktyp t
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 H0: forallx : AtomSet.elt,
x `notin` L ->
LC term ktyp (open term ktrm (tm_var (Fr x)) t)
LC term ktyp (tm_abs τ1 t)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom H0: LC term ktyp (open term ktrm (tm_var (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC term ktyp (tm_abs τ1 t)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: LC term ktyp (open term ktrm (tm_var (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC term ktyp (tm_abs τ1 t)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: LC term ktyp (open term ktrm (tm_var (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC typ ktyp τ1 /\ LC term ktyp t
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: LC term ktyp (open term ktrm (tm_var (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC typ ktyp τ1
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: LC term ktyp (open term ktrm (tm_var (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC term ktyp t
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: LC term ktyp (open term ktrm (tm_var (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC typ ktyp τ1
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: ok_type Δ τ1 e: atom IH: LC term ktyp (open term ktrm (tm_var (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC typ ktyp τ1
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: ok_type Δ τ1 e: atom IH: LC term ktyp (open term ktrm (tm_var (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
ok_type ?Δ τ1
eassumption.
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: LC term ktyp (open term ktrm (tm_var (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC term ktyp t
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: LC term ktyp
(open term ktrm (mret SystemF ktrm (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC term ktyp t
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: LC term ktyp t Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC term ktyp t
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: LC term ktyp
(open term ktrm (mret SystemF ktrm (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
ktyp <> ktrm
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: LC term ktyp t Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC term ktyp t
assumption.
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: LC term ktyp
(open term ktrm (mret SystemF ktrm (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
ktyp <> ktrm
discriminate.
Δ: kind_ctx Γ: type_ctx t1, t2: term LN τ1, τ2: typ LN j1: Δ; Γ ⊢ t1 : ty_ar τ1 τ2 j2: Δ; Γ ⊢ t2 : τ1 IHj1: LC term ktyp t1 IHj2: LC term ktyp t2
LC term ktyp (tm_app t1 t2)
now simplify_LC.
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ H0: forallx : AtomSet.elt,
x `notin` L ->
LC term ktyp (open term ktyp (ty_v (Fr x)) t)
LC term ktyp (tm_tab t)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ IH: forallx : AtomSet.elt,
x `notin` L ->
LC term ktyp (open term ktyp (ty_v (Fr x)) t)
LC term ktyp (tm_tab t)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN e0: atom H: Δ ++ e0 ~ tt; Γ ⊢ open term ktyp (ty_v (Fr e0)) t
: open typ ktyp (ty_v (Fr e0)) τ e: atom IH: LC term ktyp (open term ktyp (ty_v (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LC term ktyp (tm_tab t)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN e0: atom H: Δ ++ e0 ~ tt; Γ ⊢ open term ktyp (ty_v (Fr e0)) t
: open typ ktyp (ty_v (Fr e0)) τ e: atom IH: LC term ktyp (open term ktyp (ty_v (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LCn term ktyp 1 t
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN e0: atom H: Δ ++ e0 ~ tt; Γ ⊢ open term ktyp (ty_v (Fr e0)) t
: open typ ktyp (ty_v (Fr e0)) τ e: atom IH: LC term ktyp (open term ktyp (ty_v (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LCn term ktyp (1 - 1)
(open term ktyp (mret SystemF ktyp (Fr ?Goal1)) t)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN e0: atom H: Δ ++ e0 ~ tt; Γ ⊢ open term ktyp (ty_v (Fr e0)) t
: open typ ktyp (ty_v (Fr e0)) τ e: atom IH: LC term ktyp (open term ktyp (ty_v (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
1 > 0
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN e0: atom H: Δ ++ e0 ~ tt; Γ ⊢ open term ktyp (ty_v (Fr e0)) t
: open typ ktyp (ty_v (Fr e0)) τ e: atom IH: LC term ktyp (open term ktyp (ty_v (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
LCn term ktyp (1 - 1)
(open term ktyp (mret SystemF ktyp (Fr ?Goal1)) t)
exact IH.
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN e0: atom H: Δ ++ e0 ~ tt; Γ ⊢ open term ktyp (ty_v (Fr e0)) t
: open typ ktyp (ty_v (Fr e0)) τ e: atom IH: LC term ktyp (open term ktyp (ty_v (Fr e)) t) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
1 > 0
lia.
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: LC term ktyp t
LC term ktyp (tm_tap t τ1)
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: LC term ktyp t
LC term ktyp t
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: LC term ktyp t
LC typ ktyp τ1
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: LC term ktyp t
LC term ktyp t
assumption.
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: LC term ktyp t
LC typ ktyp τ1
nowinversion H.Qed.(** *** <<t>> is always well-scoped w.r.t. type variables *)(******************************************************************************)
forall (Δ : kind_ctx) (Γ : type_ctx) (t : term LN)
(τ : typ LN),
Δ; Γ ⊢ t : τ -> scoped term ktyp t (domset Δ)
forall (Δ : kind_ctx) (Γ : type_ctx) (t : term LN)
(τ : typ LN),
Δ; Γ ⊢ t : τ -> scoped term ktyp t (domset Δ)
Δ: kind_ctx Γ: type_ctx t: term LN τ: typ LN j: Δ; Γ ⊢ t : τ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 IH: forallx : AtomSet.elt,
x `notin` L ->
scoped term ktyp
(open term ktrm (tm_var (Fr x)) t) (domset Δ)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ IH: forallx : AtomSet.elt,
x `notin` L ->
scoped term ktyp (open term ktyp (ty_v (Fr x)) t)
(domset (Δ ++ x ~ tt))
scoped term ktyp (tm_tab t) (domset Δ)
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: scoped term ktyp t (domset Δ)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 IH: forallx : AtomSet.elt,
x `notin` L ->
scoped term ktyp
(open term ktrm (tm_var (Fr x)) t)
(domset Δ)
scoped term ktyp (tm_abs τ1 t) (domset Δ)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 IH: forallx : AtomSet.elt,
x `notin` L ->
scoped term ktyp
(open term ktrm (tm_var (Fr x)) t)
(domset Δ)
FV term ktyp (tm_abs τ1 t) ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: scoped term ktyp
(open term ktrm (tm_var (Fr e)) t)
(domset Δ) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
FV term ktyp (tm_abs τ1 t) ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: scoped term ktyp
(open term ktrm (tm_var (Fr e)) t)
(domset Δ) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
FV typ ktyp τ1 ∪ FV term ktyp t ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: scoped term ktyp
(open term ktrm (tm_var (Fr e)) t)
(domset Δ) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
FV typ ktyp τ1 ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: scoped term ktyp
(open term ktrm (tm_var (Fr e)) t)
(domset Δ) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}}) H0: FV typ ktyp τ1 ⊆ domset Δ
FV typ ktyp τ1 ∪ FV term ktyp t ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: scoped term ktyp
(open term ktrm (tm_var (Fr e)) t)
(domset Δ) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
FV typ ktyp τ1 ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: ok_type Δ τ1 e: atom IH: scoped term ktyp
(open term ktrm (tm_var (Fr e)) t)
(domset Δ) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
FV typ ktyp τ1 ⊆ domset Δ
apply H.
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: scoped term ktyp
(open term ktrm (tm_var (Fr e)) t)
(domset Δ) Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}}) H0: FV typ ktyp τ1 ⊆ domset Δ
FV typ ktyp τ1 ∪ FV term ktyp t ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: FV term ktyp (open term ktrm (tm_var (Fr e)) t)
⊆ domset Δ Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}}) H0: FV typ ktyp τ1 ⊆ domset Δ
FV typ ktyp τ1 ∪ FV term ktyp t ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN e0: atom H: Δ; Γ ++ e0 ~ τ1
⊢ open term ktrm (tm_var (Fr e0)) t : τ2 e: atom IH: FV term ktyp t ⊆ domset Δ Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}}) H0: FV typ ktyp τ1 ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ IH: forallx : AtomSet.elt,
x `notin` L ->
scoped term ktyp (open term ktyp (ty_v (Fr x)) t)
(domset (Δ ++ x ~ tt))
scoped term ktyp (tm_tab t) (domset Δ)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ IH: forallx : AtomSet.elt,
x `notin` L ->
scoped term ktyp (open term ktyp (ty_v (Fr x)) t)
(domset (Δ ++ x ~ tt))
FV term ktyp (tm_tab t) ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ IH: forallx : AtomSet.elt,
x `notin` L ->
scoped term ktyp (open term ktyp (ty_v (Fr x)) t)
(domset (Δ ++ x ~ tt))
FV term ktyp t ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ IH: forallx : AtomSet.elt,
x `notin` L ->
scoped term ktyp (open term ktyp (ty_v (Fr x)) t)
(domset (Δ ++ x ~ tt)) x: atom Hfresh: x `notin` (L ∪ FV term ktyp t)
FV term ktyp t ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ x: atom IH: scoped term ktyp (open term ktyp (ty_v (Fr x)) t)
(domset (Δ ++ x ~ tt)) Hfresh: x `notin` (L ∪ FV term ktyp t)
FV term ktyp t ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ x: atom IH: FV term ktyp (open term ktyp (ty_v (Fr x)) t)
⊆ domset (Δ ++ x ~ tt) Hfresh: x `notin` (L ∪ FV term ktyp t)
FV term ktyp t ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ x: atom IH: FV term ktyp t ⊆ domset (Δ ++ x ~ tt) Hfresh: x `notin` (L ∪ FV term ktyp t)
FV term ktyp t ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ x: atom IH: FV term ktyp t ⊆ domset (Δ ++ x ~ tt) Hfresh: x `notin` (L ∪ FV term ktyp t)
scoped term ktyp t (domset (Δ ++ ?x ~ ?x'))
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ x: atom IH: FV term ktyp t ⊆ domset (Δ ++ x ~ tt) Hfresh: x `notin` (L ∪ FV term ktyp t)
?x `notin` FV term ktyp t
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ x: atom IH: FV term ktyp t ⊆ domset (Δ ++ x ~ tt) Hfresh: x `notin` (L ∪ FV term ktyp t)
scoped term ktyp t (domset (Δ ++ ?x ~ ?x'))
eauto.
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ x: atom IH: FV term ktyp t ⊆ domset (Δ ++ x ~ tt) Hfresh: x `notin` (L ∪ FV term ktyp t)
x `notin` FV term ktyp t
fsetdec.
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: scoped term ktyp t (domset Δ)
scoped term ktyp (tm_tap t τ1) (domset Δ)
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: FV term ktyp t ⊆ domset Δ
FV term ktyp (tm_tap t τ1) ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: FV term ktyp t ⊆ domset Δ
FV term ktyp t ∪ FV typ ktyp τ1 ⊆ domset Δ
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: FV term ktyp t ⊆ domset Δ H0: FV typ ktyp τ1 ⊆ domset Δ
FV term ktyp t ∪ FV typ ktyp τ1 ⊆ domset Δ
fsetdec.Qed.(** *** <<t>> is always well-scoped w.r.t. term variables *)(******************************************************************************)
forall (Δ : kind_ctx) (Γ : type_ctx) (t : term LN)
(τ : typ LN),
Δ; Γ ⊢ t : τ -> scoped term ktrm t (domset Γ)
forall (Δ : kind_ctx) (Γ : type_ctx) (t : term LN)
(τ : typ LN),
Δ; Γ ⊢ t : τ -> scoped term ktrm t (domset Γ)
Δ: kind_ctx Γ: type_ctx t: term LN τ: typ LN j: Δ; Γ ⊢ t : τ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 IH: forallx : AtomSet.elt,
x `notin` L ->
FV term ktrm (open term ktrm (tm_var (Fr x)) t)
⊆ domset (Γ ++ x ~ τ1)
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ IH: forallx : AtomSet.elt,
x `notin` L ->
FV term ktrm (open term ktyp (ty_v (Fr x)) t)
⊆ domset Γ
FV term ktrm (tm_tab t) ⊆ domset Γ
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: FV term ktrm t ⊆ domset Γ
Δ: kind_ctx Γ: type_ctx x: atom τ: typ LN H: ok_kind_ctx Δ IH: ok_type_ctx Δ Γ Hin: x `in` domset Γ
{{x}} ⊆ domset Γ
fsetdec.
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 IH: forallx : AtomSet.elt,
x `notin` L ->
FV term ktrm (open term ktrm (tm_var (Fr x)) t)
⊆ domset (Γ ++ x ~ τ1)
FV term ktrm (tm_abs τ1 t) ⊆ domset Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 IH: forallx : AtomSet.elt,
x `notin` L ->
FV term ktrm (open term ktrm (tm_var (Fr x)) t)
⊆ domset (Γ ++ x ~ τ1)
FV typ ktrm τ1 ∪ FV term ktrm t ⊆ domset Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 IH: forallx : AtomSet.elt,
x `notin` L ->
FV term ktrm (open term ktrm (tm_var (Fr x)) t)
⊆ domset (Γ ++ x ~ τ1)
FV term ktrm t ⊆ domset Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 IH: forallx : AtomSet.elt,
x `notin` L ->
FV term ktrm (open term ktrm (tm_var (Fr x)) t)
⊆ domset (Γ ++ x ~ τ1) H0: FV term ktrm t ⊆ domset Γ
FV typ ktrm τ1 ∪ FV term ktrm t ⊆ domset Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 IH: forallx : AtomSet.elt,
x `notin` L ->
FV term ktrm (open term ktrm (tm_var (Fr x)) t)
⊆ domset (Γ ++ x ~ τ1)
FV term ktrm t ⊆ domset Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 IH: forallx : AtomSet.elt,
x `notin` L ->
FV term ktrm (open term ktrm (tm_var (Fr x)) t)
⊆ domset (Γ ++ x ~ τ1) x: atom Hfresh: x `notin` (L ∪ FV term ktrm t)
FV term ktrm t ⊆ domset Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 x: atom IH: FV term ktrm (open term ktrm (tm_var (Fr x)) t)
⊆ domset (Γ ++ x ~ τ1) Hfresh: x `notin` (L ∪ FV term ktrm t)
FV term ktrm t ⊆ domset Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 x: atom IH: FV term ktrm t ⊆ domset (Γ ++ x ~ τ1) Hfresh: x `notin` (L ∪ FV term ktrm t)
FV term ktrm t ⊆ domset Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 x: atom IH: FV term ktrm t ⊆ domset (Γ ++ x ~ τ1) Hfresh: x `notin` (L ∪ FV term ktrm t)
scoped term ktrm t (domset (Γ ++ ?x ~ ?x'))
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 x: atom IH: FV term ktrm t ⊆ domset (Γ ++ x ~ τ1) Hfresh: x `notin` (L ∪ FV term ktrm t)
?x `notin` FV term ktrm t
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 x: atom IH: FV term ktrm t ⊆ domset (Γ ++ x ~ τ1) Hfresh: x `notin` (L ∪ FV term ktrm t)
scoped term ktrm t (domset (Γ ++ ?x ~ ?x'))
eauto.
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 x: atom IH: FV term ktrm t ⊆ domset (Γ ++ x ~ τ1) Hfresh: x `notin` (L ∪ FV term ktrm t)
x `notin` FV term ktrm t
fsetdec.
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 IH: forallx : AtomSet.elt,
x `notin` L ->
FV term ktrm (open term ktrm (tm_var (Fr x)) t)
⊆ domset (Γ ++ x ~ τ1) H0: FV term ktrm t ⊆ domset Γ
FV typ ktrm τ1 ∪ FV term ktrm t ⊆ domset Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ; Γ ++ x ~ τ1 ⊢ open term ktrm (tm_var (Fr x)) t
: τ2 IH: forallx : AtomSet.elt,
x `notin` L ->
FV term ktrm (open term ktrm (tm_var (Fr x)) t)
⊆ domset (Γ ++ x ~ τ1) H0: FV term ktrm t ⊆ domset Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ IH: forallx : AtomSet.elt,
x `notin` L ->
FV term ktrm (open term ktyp (ty_v (Fr x)) t)
⊆ domset Γ
FV term ktrm (tm_tab t) ⊆ domset Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
Δ ++ x ~ tt; Γ ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ IH: forallx : AtomSet.elt,
x `notin` L ->
FV term ktrm (open term ktyp (ty_v (Fr x)) t)
⊆ domset Γ
FV term ktrm t ⊆ domset Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN e0: atom H: Δ ++ e0 ~ tt; Γ ⊢ open term ktyp (ty_v (Fr e0)) t
: open typ ktyp (ty_v (Fr e0)) τ e: atom IH: FV term ktrm (open term ktyp (ty_v (Fr e)) t)
⊆ domset Γ Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
FV term ktrm t ⊆ domset Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN e0: atom H: Δ ++ e0 ~ tt; Γ ⊢ open term ktyp (ty_v (Fr e0)) t
: open typ ktyp (ty_v (Fr e0)) τ e: atom IH: FV term ktrm (open term ktyp (ty_v (Fr e)) t)
⊆ domset Γ Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
FV term ktrm t ⊆ ?y
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN e0: atom H: Δ ++ e0 ~ tt; Γ ⊢ open term ktyp (ty_v (Fr e0)) t
: open typ ktyp (ty_v (Fr e0)) τ e: atom IH: FV term ktrm (open term ktyp (ty_v (Fr e)) t)
⊆ domset Γ Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
?y ⊆ domset Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN e0: atom H: Δ ++ e0 ~ tt; Γ ⊢ open term ktyp (ty_v (Fr e0)) t
: open typ ktyp (ty_v (Fr e0)) τ e: atom IH: FV term ktrm (open term ktyp (ty_v (Fr e)) t)
⊆ domset Γ Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}}) lemma:= FV_open_lower term t (ktyp : K) (ktrm : K): forallu : SystemF (ktyp : K) LN,
FV term (ktrm : K) t
⊆ FV term (ktrm : K)
(open term (ktyp : K) u t)
FV term ktrm t ⊆ ?y
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN e0: atom H: Δ ++ e0 ~ tt; Γ ⊢ open term ktyp (ty_v (Fr e0)) t
: open typ ktyp (ty_v (Fr e0)) τ e: atom IH: FV term ktrm (open term ktyp (ty_v (Fr e)) t)
⊆ domset Γ Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
?y ⊆ domset Γ
Δ: kind_ctx Γ: type_ctx L: AtomSet.t τ: typ LN t: term LN e0: atom H: Δ ++ e0 ~ tt; Γ ⊢ open term ktyp (ty_v (Fr e0)) t
: open typ ktyp (ty_v (Fr e0)) τ e: atom IH: FV term ktrm (open term ktyp (ty_v (Fr e)) t)
⊆ domset Γ Hfresh: e `notin` L Hfresh0: e0 `notin` (L ∪ {{e}})
FV term ktrm (open term ktyp ?u t) ⊆ domset Γ
eauto.
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: FV term ktrm t ⊆ domset Γ
FV term ktrm (tm_tap t τ1) ⊆ domset Γ
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: FV term ktrm t ⊆ domset Γ
FV term ktrm t ∪ FV typ ktrm τ1 ⊆ domset Γ
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: FV term ktrm t ⊆ domset Γ H0: FV typ ktrm τ1 ⊆ domset Γ
FV term ktrm t ∪ FV typ ktrm τ1 ⊆ domset Γ
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: FV term ktrm t ⊆ domset Γ
FV typ ktrm τ1 ⊆ domset Γ
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: FV term ktrm t ⊆ domset Γ H0: FV typ ktrm τ1 ⊆ domset Γ
FV term ktrm t ∪ FV typ ktrm τ1 ⊆ domset Γ
fsetdec.
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: FV term ktrm t ⊆ domset Γ
FV typ ktrm τ1 ⊆ domset Γ
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: FV term ktrm t ⊆ domset Γ H0: scoped typ ktyp τ1 (domset Δ) H1: LC typ ktyp τ1
FV typ ktrm τ1 ⊆ domset Γ
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: FV term ktrm t ⊆ domset Γ H0: FV typ ktyp τ1 ⊆ domset Δ H1: LC typ ktyp τ1
FV typ ktrm τ1 ⊆ domset Γ
Δ: kind_ctx Γ: type_ctx t: term LN τ1, τ2: typ LN H: ok_type Δ τ1 j: Δ; Γ ⊢ t : ty_univ τ2 IHj: FV term ktrm t ⊆ domset Γ H0: FV typ ktyp τ1 ⊆ domset Δ H1: LC typ ktyp τ1
∅ ⊆ domset Γ
fsetdec.Qed.(** *** <<t>> is always well-formed *)(******************************************************************************)
forall (Δ : kind_ctx) (Γ : type_ctx) (t : term LN)
(τ : typ LN), Δ; Γ ⊢ t : τ -> ok_term Δ Γ t
forall (Δ : kind_ctx) (Γ : type_ctx) (t : term LN)
(τ : typ LN), Δ; Γ ⊢ t : τ -> ok_term Δ Γ t
Δ: kind_ctx Γ: type_ctx t: term LN τ: typ LN H: Δ; Γ ⊢ t : τ
ok_term Δ Γ t
Δ: kind_ctx Γ: type_ctx t: term LN τ: typ LN H: Δ; Γ ⊢ t : τ
scoped term ktyp t (domset Δ) /\
scoped term ktrm t (domset Γ) /\
LC term ktrm t /\ LC term ktyp t
intuition (eauto using j_sc_KTerm_term, j_sc_ktyp_term, j_lc_KTerm_term, j_lc_ktyp_term).Qed.(** *** Substitution *)(******************************************************************************)
forall (Δ : kind_ctx) (Γ1Γ2 : list (atom * typ LN))
x (τ2 : typ LN) (tu : term LN) (τ1 : typ LN),
Δ; Γ1 ++ x ~ τ2 ++ Γ2 ⊢ t : τ1 ->
Δ; Γ1 ++ Γ2 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ2 ⊢ subst term ktrm x u t : τ1
forall (Δ : kind_ctx) (Γ1Γ2 : list (atom * typ LN))
x (τ2 : typ LN) (tu : term LN) (τ1 : typ LN),
Δ; Γ1 ++ x ~ τ2 ++ Γ2 ⊢ t : τ1 ->
Δ; Γ1 ++ Γ2 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ2 ⊢ subst term ktrm x u t : τ1
Δ: kind_ctx Γ1, Γ2: list (atom * typ LN) x: atom τ2: typ LN t, u: term LN τ1: typ LN jt: Δ; Γ1 ++ x ~ τ2 ++ Γ2 ⊢ t : τ1 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
Δ; Γ1 ++ Γ2 ⊢ subst term ktrm x u t : τ1
Δ: kind_ctx Γ1, Γ2: list (atom * typ LN) x: atom τ2: typ LN t, u: term LN τ1: typ LN Γ: list (atom * typ LN) HeqΓ: Γ = Γ1 ++ x ~ τ2 ++ Γ2 jt: Δ; Γ ⊢ t : τ1 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
Δ; Γ1 ++ Γ2 ⊢ subst term ktrm x u t : τ1
Δ: kind_ctx Γ1: list (atom * typ LN) x: atom τ2: typ LN t, u: term LN τ1: typ LN Γ: list (atom * typ LN) jt: Δ; Γ ⊢ t : τ1
forallΓ2 : list (atom * typ LN),
Γ = Γ1 ++ x ~ τ2 ++ Γ2 ->
Δ; Γ1 ++ Γ2 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ2 ⊢ subst term ktrm x u t : τ1
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) H1: (x0, τ) ∈ (Γ1 ++ x ~ τ2 ++ Γ2) H0: ok_type_ctx Δ (Γ1 ++ x ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
Δ; Γ1 ++ Γ2 ⊢ subst term ktrm x u (tm_var (Fr x0)) : τ
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) H0: forallx0 : AtomSet.elt,
x0 `notin` L ->
forallΓ3 : list (atom * typ LN),
(Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1 =
Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ subst term ktrm x u
(open term ktrm (tm_var (Fr x0)) t) : τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
Δ; Γ1 ++ Γ2 ⊢ subst term ktrm x u (tm_abs τ1 t)
: ty_ar τ1 τ0
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx t1, t2: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) IHjt2: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3 ⊢ subst term ktrm x u t2 : τ1 IHjt1: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3 ⊢ subst term ktrm x u t1
: ty_ar τ1 τ0 jt2: Δ; Γ1 ++ x ~ τ2 ++ Γ2 ⊢ t2 : τ1 jt1: Δ; Γ1 ++ x ~ τ2 ++ Γ2 ⊢ t1 : ty_ar τ1 τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
Δ; Γ1 ++ Γ2 ⊢ subst term ktrm x u (tm_app t1 t2) : τ0
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t τ: typ LN t: term LN Γ2: list (atom * typ LN) H0: forallx0 : AtomSet.elt,
x0 `notin` L ->
forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ ++ x0 ~ tt; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ ++ x0 ~ tt; Γ1 ++ Γ3
⊢ subst term ktrm x u
(open term ktyp (ty_v (Fr x0)) t)
: open typ ktyp (ty_v (Fr x0)) τ H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ ++ x0 ~ tt; Γ1 ++ x ~ τ2 ++ Γ2
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
Δ; Γ1 ++ Γ2 ⊢ subst term ktrm x u (tm_tab t)
: ty_univ τ
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx t: term LN τ1, τ0: typ LN H: ok_type Δ τ1 Γ2: list (atom * typ LN) IHjt: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3 ⊢ subst term ktrm x u t
: ty_univ τ0 jt: Δ; Γ1 ++ x ~ τ2 ++ Γ2 ⊢ t : ty_univ τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
Δ; Γ1 ++ Γ2 ⊢ subst term ktrm x u (tm_tap t τ1)
: open typ ktyp τ1 τ0
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) H1: (x0, τ) ∈ (Γ1 ++ x ~ τ2 ++ Γ2) H0: ok_type_ctx Δ (Γ1 ++ x ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
Δ; Γ1 ++ Γ2 ⊢ subst term ktrm x u (tm_var (Fr x0)) : τ
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) H1: (x0, τ) ∈ (Γ1 ++ x ~ τ2 ++ Γ2) Hok: ok_type_ctx Δ (Γ1 ++ x ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
Δ; Γ1 ++ Γ2 ⊢ subst term ktrm x u (tm_var (Fr x0)) : τ
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hvar: (x0, τ) ∈ (Γ1 ++ x ~ τ2 ++ Γ2) Hok: ok_type_ctx Δ (Γ1 ++ x ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
Δ; Γ1 ++ Γ2 ⊢ subst term ktrm x u (tm_var (Fr x0)) : τ
Γ1: list (atom * typ LN) τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hok: ok_type_ctx Δ (Γ1 ++ x0 ~ τ2 ++ Γ2) Hvar: (x0, τ) ∈ (Γ1 ++ x0 ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_EQs: x0 = x0
Δ; Γ1 ++ Γ2 ⊢ subst term ktrm x0 u (tm_var (Fr x0))
: τ
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hvar: (x0, τ) ∈ (Γ1 ++ x ~ τ2 ++ Γ2) Hok: ok_type_ctx Δ (Γ1 ++ x ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_NEQ: x <> x0 DESTR_NEQs: x0 <> x
Δ; Γ1 ++ Γ2 ⊢ subst term ktrm x u (tm_var (Fr x0)) : τ
Γ1: list (atom * typ LN) τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hok: ok_type_ctx Δ (Γ1 ++ x0 ~ τ2 ++ Γ2) Hvar: (x0, τ) ∈ (Γ1 ++ x0 ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_EQs: x0 = x0
Δ; Γ1 ++ Γ2 ⊢ subst term ktrm x0 u (tm_var (Fr x0))
: τ
Γ1: list (atom * typ LN) τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hok: ok_type_ctx Δ (Γ1 ++ x0 ~ τ2 ++ Γ2) Hvar: (x0, τ) ∈ (Γ1 ++ x0 ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_EQs: x0 = x0
Δ; Γ1 ++ Γ2 ⊢ u : τ
Γ1: list (atom * typ LN) τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hok: ok_type_ctx Δ (Γ1 ++ x0 ~ τ2 ++ Γ2) Hvar: (x0, τ) ∈ (Γ1 ++ x0 ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_EQs: x0 = x0
u = subst term ktrm x0 u (tm_var (Fr x0))
Γ1: list (atom * typ LN) τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hok: ok_type_ctx Δ (Γ1 ++ x0 ~ τ2 ++ Γ2) Hvar: (x0, τ) ∈ (Γ1 ++ x0 ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_EQs: x0 = x0
u = subst term ktrm x0 u (tm_var (Fr x0))
Γ1: list (atom * typ LN) τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hok: ok_type_ctx Δ (Γ1 ++ x0 ~ τ2 ++ Γ2) Hvar: (x0, τ) ∈ (Γ1 ++ x0 ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_EQs: x0 = x0
u = btg ktrm (subst_loc ktrm x0 u) ktrm (Fr x0)
Γ1: list (atom * typ LN) τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hok: ok_type_ctx Δ (Γ1 ++ x0 ~ τ2 ++ Γ2) Hvar: (x0, τ) ∈ (Γ1 ++ x0 ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_EQs: x0 = x0
u = subst_loc ktrm x0 u (Fr x0)
Γ1: list (atom * typ LN) τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hok: ok_type_ctx Δ (Γ1 ++ x0 ~ τ2 ++ Γ2) Hvar: (x0, τ) ∈ (Γ1 ++ x0 ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_EQs: x0 = x0
u = (if x0 == x0 then u else tm_var (Fr x0))
destruct_eq_args x0 x0.
Γ1: list (atom * typ LN) τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hok: ok_type_ctx Δ (Γ1 ++ x0 ~ τ2 ++ Γ2) Hvar: (x0, τ) ∈ (Γ1 ++ x0 ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_EQs: x0 = x0
Δ; Γ1 ++ Γ2 ⊢ u : τ
Γ1: list (atom * typ LN) τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hok: ok_type_ctx Δ (Γ1 ++ x0 ~ τ2 ++ Γ2) Hvar: (x0, τ) ∈ (Γ1 ++ x0 ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_EQs: x0 = x0
τ = τ2
Γ1: list (atom * typ LN) τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hok: ok_type_ctx Δ (Γ1 ++ x0 ~ τ2 ++ Γ2) Hvar: (x0, τ) ∈ (Γ1 ++ x0 ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_EQs: x0 = x0
(?x, τ) ∈ (?Γ1 ++ ?x ~ τ2 ++ ?Γ2)
Γ1: list (atom * typ LN) τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hok: ok_type_ctx Δ (Γ1 ++ x0 ~ τ2 ++ Γ2) Hvar: (x0, τ) ∈ (Γ1 ++ x0 ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_EQs: x0 = x0
uniq (?Γ1 ++ ?x ~ τ2 ++ ?Γ2)
Γ1: list (atom * typ LN) τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hok: ok_type_ctx Δ (Γ1 ++ x0 ~ τ2 ++ Γ2) Hvar: (x0, τ) ∈ (Γ1 ++ x0 ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_EQs: x0 = x0
uniq (Γ1 ++ x0 ~ τ2 ++ Γ2)
apply Hok.
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hvar: (x0, τ) ∈ (Γ1 ++ x ~ τ2 ++ Γ2) Hok: ok_type_ctx Δ (Γ1 ++ x ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_NEQ: x <> x0 DESTR_NEQs: x0 <> x
Δ; Γ1 ++ Γ2 ⊢ subst term ktrm x u (tm_var (Fr x0)) : τ
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hvar: (x0, τ) ∈ (Γ1 ++ x ~ τ2 ++ Γ2) Hok: ok_type_ctx Δ (Γ1 ++ x ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_NEQ: x <> x0 DESTR_NEQs: x0 <> x
Δ; Γ1 ++ Γ2 ⊢ tm_var (Fr x0) : τ
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hvar: (x0, τ) ∈ (Γ1 ++ x ~ τ2 ++ Γ2) Hok: ok_type_ctx Δ (Γ1 ++ x ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_NEQ: x <> x0 DESTR_NEQs: x0 <> x
tm_var (Fr x0) = subst term ktrm x u (tm_var (Fr x0))
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hvar: (x0, τ) ∈ (Γ1 ++ x ~ τ2 ++ Γ2) Hok: ok_type_ctx Δ (Γ1 ++ x ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_NEQ: x <> x0 DESTR_NEQs: x0 <> x
tm_var (Fr x0) = subst term ktrm x u (tm_var (Fr x0))
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hvar: (x0, τ) ∈ (Γ1 ++ x ~ τ2 ++ Γ2) Hok: ok_type_ctx Δ (Γ1 ++ x ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_NEQ: x <> x0 DESTR_NEQs: x0 <> x
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hvar: (x0, τ) ∈ (Γ1 ++ x ~ τ2 ++ Γ2) Hok: ok_type_ctx Δ (Γ1 ++ x ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_NEQ: x <> x0 DESTR_NEQs: x0 <> x
tm_var (Fr x0) = subst_loc ktrm x u (Fr x0)
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hvar: (x0, τ) ∈ (Γ1 ++ x ~ τ2 ++ Γ2) Hok: ok_type_ctx Δ (Γ1 ++ x ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_NEQ: x <> x0 DESTR_NEQs: x0 <> x
tm_var (Fr x0) =
(if x == x0 then u else tm_var (Fr x0))
destruct_eq_args x x0.
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hvar: (x0, τ) ∈ (Γ1 ++ x ~ τ2 ++ Γ2) Hok: ok_type_ctx Δ (Γ1 ++ x ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_NEQ: x <> x0 DESTR_NEQs: x0 <> x
Δ; Γ1 ++ Γ2 ⊢ tm_var (Fr x0) : τ
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hvar: (x0, τ) ∈ (Γ1 ++ x ~ τ2 ++ Γ2) Hok: ok_type_ctx Δ (Γ1 ++ x ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_NEQ: x <> x0 DESTR_NEQs: x0 <> x
ok_kind_ctx Δ
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hvar: (x0, τ) ∈ (Γ1 ++ x ~ τ2 ++ Γ2) Hok: ok_type_ctx Δ (Γ1 ++ x ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_NEQ: x <> x0 DESTR_NEQs: x0 <> x
ok_type_ctx Δ (Γ1 ++ Γ2)
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hvar: (x0, τ) ∈ (Γ1 ++ x ~ τ2 ++ Γ2) Hok: ok_type_ctx Δ (Γ1 ++ x ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_NEQ: x <> x0 DESTR_NEQs: x0 <> x
(x0, τ) ∈ (Γ1 ++ Γ2)
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hvar: (x0, τ) ∈ (Γ1 ++ x ~ τ2 ++ Γ2) Hok: ok_type_ctx Δ (Γ1 ++ x ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_NEQ: x <> x0 DESTR_NEQs: x0 <> x
ok_kind_ctx Δ
assumption.
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hvar: (x0, τ) ∈ (Γ1 ++ x ~ τ2 ++ Γ2) Hok: ok_type_ctx Δ (Γ1 ++ x ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_NEQ: x <> x0 DESTR_NEQs: x0 <> x
ok_type_ctx Δ (Γ1 ++ Γ2)
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hvar: (x0, τ) ∈ (Γ1 ++ x ~ τ2 ++ Γ2) Hok: ok_type_ctx Δ (Γ1 ++ x ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_NEQ: x <> x0 DESTR_NEQs: x0 <> x
(x0, τ) ∈ (Γ1 ++ Γ2)
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hvar: (x0, τ) ∈ (Γ1 ++ x ~ τ2 ++ Γ2) Hok: ok_type_ctx Δ (Γ1 ++ x ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_NEQ: x <> x0 DESTR_NEQs: x0 <> x
ok_type_ctx Δ (Γ1 ++ Γ2)
eauto with sysf_ctx.
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hvar: (x0, τ) ∈ (Γ1 ++ x ~ τ2 ++ Γ2) Hok: ok_type_ctx Δ (Γ1 ++ x ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_NEQ: x <> x0 DESTR_NEQs: x0 <> x
(x0, τ) ∈ (Γ1 ++ Γ2)
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx x0: atom τ: typ LN H: ok_kind_ctx Δ Γ2: list (atom * typ LN) Hvar: (x0, τ) ∈ (Γ1 ++ x ~ τ2 ++ Γ2) Hok: ok_type_ctx Δ (Γ1 ++ x ~ τ2 ++ Γ2) ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 DESTR_NEQ: x <> x0 DESTR_NEQs: x0 <> x
(x0, τ) ∈ (Γ1 ++ Γ2)
eauto using binds_remove_mid.}
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) H0: forallx0 : AtomSet.elt,
x0 `notin` L ->
forallΓ3 : list (atom * typ LN),
(Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1 =
Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ subst term ktrm x u
(open term ktrm (tm_var (Fr x0)) t) : τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
Δ; Γ1 ++ Γ2 ⊢ subst term ktrm x u (tm_abs τ1 t)
: ty_ar τ1 τ0
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) IH: forallx0 : AtomSet.elt,
x0 `notin` L ->
forallΓ3 : list (atom * typ LN),
(Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1 =
Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ subst term ktrm x u
(open term ktrm (tm_var (Fr x0)) t) : τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
Δ; Γ1 ++ Γ2 ⊢ subst term ktrm x u (tm_abs τ1 t)
: ty_ar τ1 τ0
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) IH: forallx0 : AtomSet.elt,
x0 `notin` L ->
forallΓ3 : list (atom * typ LN),
(Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1 =
Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ subst term ktrm x u
(open term ktrm (tm_var (Fr x0)) t) : τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
Δ; Γ1 ++ Γ2
⊢ tm_abs (subst typ ktrm x u τ1)
(subst term ktrm x u t) :
ty_ar τ1 τ0
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) IH: forallx0 : AtomSet.elt,
x0 `notin` L ->
forallΓ3 : list (atom * typ LN),
(Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1 =
Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ subst term ktrm x u
(open term ktrm (tm_var (Fr x0)) t) : τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
Δ; Γ1 ++ Γ2 ⊢ tm_abs τ1 (subst term ktrm x u t)
: ty_ar τ1 τ0
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) IH: forallx0 : AtomSet.elt,
x0 `notin` L ->
forallΓ3 : list (atom * typ LN),
(Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1 =
Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ subst term ktrm x u
(open term ktrm (tm_var (Fr x0)) t) : τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
forallx0 : AtomSet.elt,
x0 `notin` (L ∪ {{x}} ∪ domset Γ1 ∪ domset Γ2) ->
Δ; (Γ1 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0))
(subst term ktrm x u t) : τ0
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
(Γ1 ++ x ~ τ2 ++ Γ2) ++ e ~ τ1 =
Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ subst term ktrm x u
(open term ktrm (tm_var (Fr e)) t) : τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ {{x}} ∪ domset Γ1 ∪ domset Γ2)
Δ; (Γ1 ++ Γ2) ++ e ~ τ1
⊢ open term ktrm (tm_var (Fr e))
(subst term ktrm x u t) : τ0
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
(Γ1 ++ x ~ τ2 ++ Γ2) ++ e ~ τ1 =
Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ open term ktrm
(subst (SystemF ktrm) ktrm x u
(tm_var (Fr e))) (subst term ktrm x u t)
: τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ {{x}} ∪ domset Γ1 ∪ domset Γ2)
Δ; (Γ1 ++ Γ2) ++ e ~ τ1
⊢ open term ktrm (tm_var (Fr e))
(subst term ktrm x u t) : τ0
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
(Γ1 ++ x ~ τ2 ++ Γ2) ++ e ~ τ1 =
Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ subst term ktrm x u
(open term ktrm (tm_var (Fr e)) t) : τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ {{x}} ∪ domset Γ1 ∪ domset Γ2)
LC (SystemF ktrm) ktrm u
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
(Γ1 ++ x ~ τ2 ++ Γ2) ++ e ~ τ1 =
Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ subst term ktrm x u
(open term ktrm (tm_var (Fr e)) t) : τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ {{x}} ∪ domset Γ1 ∪ domset Γ2)
LC (SystemF ktrm) ktrm u
eapply j_lc_KTerm_term; eauto.
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
(Γ1 ++ x ~ τ2 ++ Γ2) ++ e ~ τ1 =
Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ open term ktrm
(subst (SystemF ktrm) ktrm x u
(tm_var (Fr e))) (subst term ktrm x u t)
: τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ {{x}} ∪ domset Γ1 ∪ domset Γ2)
Δ; (Γ1 ++ Γ2) ++ e ~ τ1
⊢ open term ktrm (tm_var (Fr e))
(subst term ktrm x u t) : τ0
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
(Γ1 ++ x ~ τ2 ++ Γ2) ++ e ~ τ1 =
Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ open term ktrm (tm_var (Fr e))
(subst term ktrm x u t) : τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ {{x}} ∪ domset Γ1 ∪ domset Γ2)
Δ; (Γ1 ++ Γ2) ++ e ~ τ1
⊢ open term ktrm (tm_var (Fr e))
(subst term ktrm x u t) : τ0
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
(Γ1 ++ x ~ τ2 ++ Γ2) ++ e ~ τ1 =
Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ open term ktrm
(subst (SystemF ktrm) ktrm x u
(tm_var (Fr e))) (subst term ktrm x u t)
: τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ {{x}} ∪ domset Γ1 ∪ domset Γ2)
~ x ∈ free term ktrm (tm_var (Fr e))
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
(Γ1 ++ x ~ τ2 ++ Γ2) ++ e ~ τ1 =
Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ open term ktrm
(subst (SystemF ktrm) ktrm x u
(tm_var (Fr e))) (subst term ktrm x u t)
: τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ {{x}} ∪ domset Γ1 ∪ domset Γ2)
~ x ∈ free term ktrm (tm_var (Fr e))
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
(Γ1 ++ x ~ τ2 ++ Γ2) ++ e ~ τ1 =
Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ open term ktrm
(subst (SystemF ktrm) ktrm x u
(tm_var (Fr e))) (subst term ktrm x u t)
: τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ {{x}} ∪ domset Γ1 ∪ domset Γ2)
~ x ∈ [e]
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
(Γ1 ++ x ~ τ2 ++ Γ2) ++ e ~ τ1 =
Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ open term ktrm
(subst (SystemF ktrm) ktrm x u
(tm_var (Fr e))) (subst term ktrm x u t)
: τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ {{x}} ∪ domset Γ1 ∪ domset Γ2)
x <> e
fsetdec.
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
(Γ1 ++ x ~ τ2 ++ Γ2) ++ e ~ τ1 =
Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ open term ktrm (tm_var (Fr e))
(subst term ktrm x u t) : τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ {{x}} ∪ domset Γ1 ∪ domset Γ2)
Δ; (Γ1 ++ Γ2) ++ e ~ τ1
⊢ open term ktrm (tm_var (Fr e))
(subst term ktrm x u t) : τ0
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 ++ e ~ τ1 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ open term ktrm (tm_var (Fr e))
(subst term ktrm x u t) : τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ {{x}} ∪ domset Γ1 ∪ domset Γ2)
Δ; Γ1 ++ Γ2 ++ e ~ τ1
⊢ open term ktrm (tm_var (Fr e))
(subst term ktrm x u t) : τ0
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 ++ e ~ τ1 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ open term ktrm (tm_var (Fr e))
(subst term ktrm x u t) : τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ {{x}} ∪ domset Γ1 ∪ domset Γ2)
Δ; Γ1 ++ Γ2 ++ e ~ τ1 ⊢ u : τ2
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 ++ e ~ τ1 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ open term ktrm (tm_var (Fr e))
(subst term ktrm x u t) : τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ {{x}} ∪ domset Γ1 ∪ domset Γ2)
Δ; (Γ1 ++ Γ2) ++ e ~ τ1 ⊢ u : τ2
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 ++ e ~ τ1 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ open term ktrm (tm_var (Fr e))
(subst term ktrm x u t) : τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ {{x}} ∪ domset Γ1 ∪ domset Γ2)
ok_type_ctx Δ (e ~ τ1)
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 ++ e ~ τ1 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ open term ktrm (tm_var (Fr e))
(subst term ktrm x u t) : τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ {{x}} ∪ domset Γ1 ∪ domset Γ2)
disjoint (Γ1 ++ Γ2) (e ~ τ1)
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 ++ e ~ τ1 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ open term ktrm (tm_var (Fr e))
(subst term ktrm x u t) : τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ {{x}} ∪ domset Γ1 ∪ domset Γ2)
Δ; Γ1 ++ Γ2 ⊢ u : τ2
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 ++ e ~ τ1 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ open term ktrm (tm_var (Fr e))
(subst term ktrm x u t) : τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ {{x}} ∪ domset Γ1 ∪ domset Γ2)
ok_type_ctx Δ (e ~ τ1)
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 ++ e ~ τ1 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ open term ktrm (tm_var (Fr e))
(subst term ktrm x u t) : τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ {{x}} ∪ domset Γ1 ∪ domset Γ2)
ok_type Δ τ1
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 ++ e ~ τ1 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ open term ktrm (tm_var (Fr e))
(subst term ktrm x u t) : τ0 H: Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ e ~ τ1
⊢ open term ktrm (tm_var (Fr e)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ {{x}} ∪ domset Γ1 ∪ domset Γ2)
ok_type Δ τ1
eapply j_type_ctx2; eassumption.
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 ++ e ~ τ1 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ open term ktrm (tm_var (Fr e))
(subst term ktrm x u t) : τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ {{x}} ∪ domset Γ1 ∪ domset Γ2)
disjoint (Γ1 ++ Γ2) (e ~ τ1)
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 ++ e ~ τ1 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ open term ktrm (tm_var (Fr e))
(subst term ktrm x u t) : τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ {{x}} ∪ domset Γ1 ∪ domset Γ2)
e `notin` domset Γ1 /\ e `notin` domset Γ2
intuition fsetdec.
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t t: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 ++ e ~ τ1 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3
⊢ open term ktrm (tm_var (Fr e))
(subst term ktrm x u t) : τ0 H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ; (Γ1 ++ x ~ τ2 ++ Γ2) ++ x0 ~ τ1
⊢ open term ktrm (tm_var (Fr x0)) t : τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ {{x}} ∪ domset Γ1 ∪ domset Γ2)
Δ; Γ1 ++ Γ2 ⊢ u : τ2
assumption.
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx t1, t2: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) IHjt2: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3 ⊢ subst term ktrm x u t2 : τ1 IHjt1: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3 ⊢ subst term ktrm x u t1
: ty_ar τ1 τ0 jt2: Δ; Γ1 ++ x ~ τ2 ++ Γ2 ⊢ t2 : τ1 jt1: Δ; Γ1 ++ x ~ τ2 ++ Γ2 ⊢ t1 : ty_ar τ1 τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
Δ; Γ1 ++ Γ2 ⊢ subst term ktrm x u (tm_app t1 t2) : τ0
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx t1, t2: term LN τ1, τ0: typ LN Γ2: list (atom * typ LN) IHjt2: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3 ⊢ subst term ktrm x u t2 : τ1 IHjt1: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3 ⊢ subst term ktrm x u t1
: ty_ar τ1 τ0 jt2: Δ; Γ1 ++ x ~ τ2 ++ Γ2 ⊢ t2 : τ1 jt1: Δ; Γ1 ++ x ~ τ2 ++ Γ2 ⊢ t1 : ty_ar τ1 τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
Δ; Γ1 ++ Γ2
⊢ tm_app (subst term ktrm x u t1)
(subst term ktrm x u t2) : τ0
eauto using j_app.
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t τ: typ LN t: term LN Γ2: list (atom * typ LN) H0: forallx0 : AtomSet.elt,
x0 `notin` L ->
forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ ++ x0 ~ tt; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ ++ x0 ~ tt; Γ1 ++ Γ3
⊢ subst term ktrm x u
(open term ktyp (ty_v (Fr x0)) t)
: open typ ktyp (ty_v (Fr x0)) τ H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ ++ x0 ~ tt; Γ1 ++ x ~ τ2 ++ Γ2
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
Δ; Γ1 ++ Γ2 ⊢ subst term ktrm x u (tm_tab t)
: ty_univ τ
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t τ: typ LN t: term LN Γ2: list (atom * typ LN) IH: forallx0 : AtomSet.elt,
x0 `notin` L ->
forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ ++ x0 ~ tt; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ ++ x0 ~ tt; Γ1 ++ Γ3
⊢ subst term ktrm x u
(open term ktyp (ty_v (Fr x0)) t)
: open typ ktyp (ty_v (Fr x0)) τ H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ ++ x0 ~ tt; Γ1 ++ x ~ τ2 ++ Γ2
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
Δ; Γ1 ++ Γ2 ⊢ subst term ktrm x u (tm_tab t)
: ty_univ τ
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t τ: typ LN t: term LN Γ2: list (atom * typ LN) IH: forallx0 : AtomSet.elt,
x0 `notin` L ->
forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ ++ x0 ~ tt; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ ++ x0 ~ tt; Γ1 ++ Γ3
⊢ subst term ktrm x u
(open term ktyp (ty_v (Fr x0)) t)
: open typ ktyp (ty_v (Fr x0)) τ H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ ++ x0 ~ tt; Γ1 ++ x ~ τ2 ++ Γ2
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
Δ; Γ1 ++ Γ2 ⊢ tm_tab (subst term ktrm x u t)
: ty_univ τ
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t τ: typ LN t: term LN Γ2: list (atom * typ LN) IH: forallx0 : AtomSet.elt,
x0 `notin` L ->
forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ ++ x0 ~ tt; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ ++ x0 ~ tt; Γ1 ++ Γ3
⊢ subst term ktrm x u
(open term ktyp (ty_v (Fr x0)) t)
: open typ ktyp (ty_v (Fr x0)) τ H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ ++ x0 ~ tt; Γ1 ++ x ~ τ2 ++ Γ2
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
forallx0 : AtomSet.elt,
x0 `notin` (L ∪ domset Δ) ->
Δ ++ x0 ~ tt; Γ1 ++ Γ2
⊢ open term ktyp (ty_v (Fr x0))
(subst term ktrm x u t)
: open typ ktyp (ty_v (Fr x0)) τ
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t τ: typ LN t: term LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ ++ e ~ tt; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ ++ e ~ tt; Γ1 ++ Γ3
⊢ subst term ktrm x u
(open term ktyp (ty_v (Fr e)) t)
: open typ ktyp (ty_v (Fr e)) τ H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ ++ x0 ~ tt; Γ1 ++ x ~ τ2 ++ Γ2
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ domset Δ)
Δ ++ e ~ tt; Γ1 ++ Γ2
⊢ open term ktyp (ty_v (Fr e)) (subst term ktrm x u t)
: open typ ktyp (ty_v (Fr e)) τ
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t τ: typ LN t: term LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ ++ e ~ tt; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ ++ e ~ tt; Γ1 ++ Γ3
⊢ open term ktyp
(subst (SystemF ktyp) ktrm x u (ty_v (Fr e)))
(subst term ktrm x u t)
: open typ ktyp (ty_v (Fr e)) τ H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ ++ x0 ~ tt; Γ1 ++ x ~ τ2 ++ Γ2
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ domset Δ)
Δ ++ e ~ tt; Γ1 ++ Γ2
⊢ open term ktyp (ty_v (Fr e)) (subst term ktrm x u t)
: open typ ktyp (ty_v (Fr e)) τ
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t τ: typ LN t: term LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ ++ e ~ tt; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ ++ e ~ tt; Γ1 ++ Γ3
⊢ subst term ktrm x u
(open term ktyp (ty_v (Fr e)) t)
: open typ ktyp (ty_v (Fr e)) τ H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ ++ x0 ~ tt; Γ1 ++ x ~ τ2 ++ Γ2
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ domset Δ)
ktyp <> ktrm
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t τ: typ LN t: term LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ ++ e ~ tt; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ ++ e ~ tt; Γ1 ++ Γ3
⊢ subst term ktrm x u
(open term ktyp (ty_v (Fr e)) t)
: open typ ktyp (ty_v (Fr e)) τ H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ ++ x0 ~ tt; Γ1 ++ x ~ τ2 ++ Γ2
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ domset Δ)
LC (SystemF ktrm) ktyp u
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t τ: typ LN t: term LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ ++ e ~ tt; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ ++ e ~ tt; Γ1 ++ Γ3
⊢ subst term ktrm x u
(open term ktyp (ty_v (Fr e)) t)
: open typ ktyp (ty_v (Fr e)) τ H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ ++ x0 ~ tt; Γ1 ++ x ~ τ2 ++ Γ2
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ domset Δ)
ktyp <> ktrm
discriminate.
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t τ: typ LN t: term LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ ++ e ~ tt; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ ++ e ~ tt; Γ1 ++ Γ3
⊢ open term ktyp
(subst (SystemF ktyp) ktrm x u (ty_v (Fr e)))
(subst term ktrm x u t)
: open typ ktyp (ty_v (Fr e)) τ H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ ++ x0 ~ tt; Γ1 ++ x ~ τ2 ++ Γ2
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ domset Δ)
Δ ++ e ~ tt; Γ1 ++ Γ2
⊢ open term ktyp (ty_v (Fr e)) (subst term ktrm x u t)
: open typ ktyp (ty_v (Fr e)) τ
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t τ: typ LN t: term LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ ++ e ~ tt; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ ++ e ~ tt; Γ1 ++ Γ3
⊢ subst term ktrm x u
(open term ktyp (ty_v (Fr e)) t)
: open typ ktyp (ty_v (Fr e)) τ H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ ++ x0 ~ tt; Γ1 ++ x ~ τ2 ++ Γ2
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ domset Δ)
LC (SystemF ktrm) ktyp u
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t τ: typ LN t: term LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ ++ e ~ tt; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ ++ e ~ tt; Γ1 ++ Γ3
⊢ subst term ktrm x u
(open term ktyp (ty_v (Fr e)) t)
: open typ ktyp (ty_v (Fr e)) τ H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ ++ x0 ~ tt; Γ1 ++ x ~ τ2 ++ Γ2
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ domset Δ)
LC (SystemF ktrm) ktyp u
eapply j_ok_term; eassumption.
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t τ: typ LN t: term LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ ++ e ~ tt; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ ++ e ~ tt; Γ1 ++ Γ3
⊢ open term ktyp
(subst (SystemF ktyp) ktrm x u (ty_v (Fr e)))
(subst term ktrm x u t)
: open typ ktyp (ty_v (Fr e)) τ H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ ++ x0 ~ tt; Γ1 ++ x ~ τ2 ++ Γ2
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ domset Δ)
Δ ++ e ~ tt; Γ1 ++ Γ2
⊢ open term ktyp (ty_v (Fr e)) (subst term ktrm x u t)
: open typ ktyp (ty_v (Fr e)) τ
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t τ: typ LN t: term LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ ++ e ~ tt; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ ++ e ~ tt; Γ1 ++ Γ3
⊢ open term ktyp
(subst (SystemF ktyp) ktrm x u (ty_v (Fr e)))
(subst term ktrm x u t)
: open typ ktyp (ty_v (Fr e)) τ H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ ++ x0 ~ tt; Γ1 ++ x ~ τ2 ++ Γ2
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ domset Δ)
Δ ++ e ~ tt; Γ1 ++ Γ2 ⊢ u : τ2
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t τ: typ LN t: term LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ ++ e ~ tt; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ ++ e ~ tt; Γ1 ++ Γ3
⊢ open term ktyp
(subst (SystemF ktyp) ktrm x u (ty_v (Fr e)))
(subst term ktrm x u t)
: open typ ktyp (ty_v (Fr e)) τ H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ ++ x0 ~ tt; Γ1 ++ x ~ τ2 ++ Γ2
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ domset Δ)
ok_kind_ctx (e ~ tt)
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t τ: typ LN t: term LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ ++ e ~ tt; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ ++ e ~ tt; Γ1 ++ Γ3
⊢ open term ktyp
(subst (SystemF ktyp) ktrm x u (ty_v (Fr e)))
(subst term ktrm x u t)
: open typ ktyp (ty_v (Fr e)) τ H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ ++ x0 ~ tt; Γ1 ++ x ~ τ2 ++ Γ2
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ domset Δ)
disjoint Δ (e ~ tt)
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t τ: typ LN t: term LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ ++ e ~ tt; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ ++ e ~ tt; Γ1 ++ Γ3
⊢ open term ktyp
(subst (SystemF ktyp) ktrm x u (ty_v (Fr e)))
(subst term ktrm x u t)
: open typ ktyp (ty_v (Fr e)) τ H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ ++ x0 ~ tt; Γ1 ++ x ~ τ2 ++ Γ2
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ domset Δ)
Δ; Γ1 ++ Γ2 ⊢ u : τ2
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t τ: typ LN t: term LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ ++ e ~ tt; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ ++ e ~ tt; Γ1 ++ Γ3
⊢ open term ktyp
(subst (SystemF ktyp) ktrm x u (ty_v (Fr e)))
(subst term ktrm x u t)
: open typ ktyp (ty_v (Fr e)) τ H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ ++ x0 ~ tt; Γ1 ++ x ~ τ2 ++ Γ2
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ domset Δ)
ok_kind_ctx (e ~ tt)
apply ok_kind_ctx_one.
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t τ: typ LN t: term LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ ++ e ~ tt; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ ++ e ~ tt; Γ1 ++ Γ3
⊢ open term ktyp
(subst (SystemF ktyp) ktrm x u (ty_v (Fr e)))
(subst term ktrm x u t)
: open typ ktyp (ty_v (Fr e)) τ H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ ++ x0 ~ tt; Γ1 ++ x ~ τ2 ++ Γ2
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ domset Δ)
disjoint Δ (e ~ tt)
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t τ: typ LN t: term LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ ++ e ~ tt; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ ++ e ~ tt; Γ1 ++ Γ3
⊢ open term ktyp
(subst (SystemF ktyp) ktrm x u (ty_v (Fr e)))
(subst term ktrm x u t)
: open typ ktyp (ty_v (Fr e)) τ H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ ++ x0 ~ tt; Γ1 ++ x ~ τ2 ++ Γ2
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ domset Δ)
e `notin` domset Δ
fsetdec.
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx L: AtomSet.t τ: typ LN t: term LN Γ2: list (atom * typ LN) e: AtomSet.elt IH: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ ++ e ~ tt; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ ++ e ~ tt; Γ1 ++ Γ3
⊢ open term ktyp
(subst (SystemF ktyp) ktrm x u (ty_v (Fr e)))
(subst term ktrm x u t)
: open typ ktyp (ty_v (Fr e)) τ H: forallx0 : AtomSet.elt,
x0 `notin` L ->
Δ ++ x0 ~ tt; Γ1 ++ x ~ τ2 ++ Γ2
⊢ open term ktyp (ty_v (Fr x0)) t
: open typ ktyp (ty_v (Fr x0)) τ ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2 Notin: e `notin` (L ∪ domset Δ)
Δ; Γ1 ++ Γ2 ⊢ u : τ2
assumption.
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx t: term LN τ1, τ0: typ LN H: ok_type Δ τ1 Γ2: list (atom * typ LN) IHjt: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3 ⊢ subst term ktrm x u t
: ty_univ τ0 jt: Δ; Γ1 ++ x ~ τ2 ++ Γ2 ⊢ t : ty_univ τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
Δ; Γ1 ++ Γ2 ⊢ subst term ktrm x u (tm_tap t τ1)
: open typ ktyp τ1 τ0
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx t: term LN τ1, τ0: typ LN H: ok_type Δ τ1 Γ2: list (atom * typ LN) IHjt: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3 ⊢ subst term ktrm x u t
: ty_univ τ0 jt: Δ; Γ1 ++ x ~ τ2 ++ Γ2 ⊢ t : ty_univ τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
Δ; Γ1 ++ Γ2
⊢ tm_tap (subst term ktrm x u t)
(subst typ ktrm x u τ1) :
open typ ktyp τ1 τ0
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx t: term LN τ1, τ0: typ LN H: ok_type Δ τ1 Γ2: list (atom * typ LN) IHjt: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3 ⊢ subst term ktrm x u t
: ty_univ τ0 jt: Δ; Γ1 ++ x ~ τ2 ++ Γ2 ⊢ t : ty_univ τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
Δ; Γ1 ++ Γ2 ⊢ tm_tap (subst term ktrm x u t) τ1
: open typ ktyp τ1 τ0
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx t: term LN τ1, τ0: typ LN H: ok_type Δ τ1 Γ2: list (atom * typ LN) IHjt: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3 ⊢ subst term ktrm x u t
: ty_univ τ0 jt: Δ; Γ1 ++ x ~ τ2 ++ Γ2 ⊢ t : ty_univ τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
ok_type Δ τ1
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx t: term LN τ1, τ0: typ LN H: ok_type Δ τ1 Γ2: list (atom * typ LN) IHjt: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3 ⊢ subst term ktrm x u t
: ty_univ τ0 jt: Δ; Γ1 ++ x ~ τ2 ++ Γ2 ⊢ t : ty_univ τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
Δ; Γ1 ++ Γ2 ⊢ subst term ktrm x u t : ty_univ τ0
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx t: term LN τ1, τ0: typ LN H: ok_type Δ τ1 Γ2: list (atom * typ LN) IHjt: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3 ⊢ subst term ktrm x u t
: ty_univ τ0 jt: Δ; Γ1 ++ x ~ τ2 ++ Γ2 ⊢ t : ty_univ τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
ok_type Δ τ1
assumption.
Γ1: list (atom * typ LN) x: atom τ2: typ LN u: term LN Δ: kind_ctx t: term LN τ1, τ0: typ LN H: ok_type Δ τ1 Γ2: list (atom * typ LN) IHjt: forallΓ3 : list (atom * typ LN),
Γ1 ++ x ~ τ2 ++ Γ2 = Γ1 ++ x ~ τ2 ++ Γ3 ->
Δ; Γ1 ++ Γ3 ⊢ u : τ2 ->
Δ; Γ1 ++ Γ3 ⊢ subst term ktrm x u t
: ty_univ τ0 jt: Δ; Γ1 ++ x ~ τ2 ++ Γ2 ⊢ t : ty_univ τ0 ju: Δ; Γ1 ++ Γ2 ⊢ u : τ2
Δ; Γ1 ++ Γ2 ⊢ subst term ktrm x u t : ty_univ τ0
auto.Qed.
forall (Δ : kind_ctx) (Γ : list (atom * typ LN)) x
(τ2 : typ LN) (tu : term LN) (τ1 : typ LN),
Δ; Γ ++ x ~ τ2 ⊢ t : τ1 ->
Δ; Γ ⊢ u : τ2 -> Δ; Γ ⊢ subst term ktrm x u t : τ1
forall (Δ : kind_ctx) (Γ : list (atom * typ LN)) x
(τ2 : typ LN) (tu : term LN) (τ1 : typ LN),
Δ; Γ ++ x ~ τ2 ⊢ t : τ1 ->
Δ; Γ ⊢ u : τ2 -> Δ; Γ ⊢ subst term ktrm x u t : τ1
Δ: kind_ctx Γ: list (atom * typ LN) x: atom τ2: typ LN t, u: term LN τ1: typ LN jt: Δ; Γ ++ x ~ τ2 ⊢ t : τ1 ju: Δ; Γ ⊢ u : τ2
Δ; Γ ⊢ subst term ktrm x u t : τ1
Δ: kind_ctx Γ: list (atom * typ LN) x: atom τ2: typ LN t, u: term LN τ1: typ LN jt: Δ; Γ ++ x ~ τ2 ++ [] ⊢ t : τ1 ju: Δ; Γ ⊢ u : τ2
Δ; Γ ⊢ subst term ktrm x u t : τ1
Δ: kind_ctx Γ: list (atom * typ LN) x: atom τ2: typ LN t, u: term LN τ1: typ LN jt: Δ; Γ ++ x ~ τ2 ++ [] ⊢ t : τ1 ju: Δ; Γ ++ [] ⊢ u : τ2
Δ; Γ ⊢ subst term ktrm x u t : τ1
Δ: kind_ctx Γ: list (atom * typ LN) x: atom τ2: typ LN t, u: term LN τ1: typ LN jt: Δ; Γ ++ x ~ τ2 ++ [] ⊢ t : τ1 ju: Δ; Γ ++ [] ⊢ u : τ2
Δ; Γ ++ [] ⊢ subst term ktrm x u t : τ1
eauto using j_type_ctx_subst.Qed.(** * Progress and preservation *)(******************************************************************************)
preservation
preservation
t, t': term LN τ: typ LN j: []; [] ⊢ t : τ
red t t' -> []; [] ⊢ t' : τ
t: term LN τ: typ LN j: []; [] ⊢ t : τ
forallt' : term LN, red t t' -> []; [] ⊢ t' : τ
t: term LN τ: typ LN Delta: list (atom * unit) HeqDelta: Delta = [] j: Delta; [] ⊢ t : τ
forallt' : term LN, red t t' -> Delta; [] ⊢ t' : τ
t: term LN τ: typ LN Delta: list (atom * unit) HeqDelta: Delta = [] Gamma: list (atom * typ LN) HeqGamma: Gamma = [] j: Delta; Gamma ⊢ t : τ
forallt' : term LN, red t t' -> Delta; Gamma ⊢ t' : τ
x: atom τ: typ LN H0: ok_type_ctx [] [] H: ok_kind_ctx [] H1: (x, τ) ∈ []
forallt' : term LN,
red (tm_var (Fr x)) t' -> []; [] ⊢ t' : τ
L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
[]; [] ++ x ~ τ1
⊢ open term ktrm (tm_var (Fr x)) t : τ2 H0: forallx : AtomSet.elt,
x `notin` L ->
[] = [] ->
[] ++ x ~ τ1 = [] ->
forallt' : term LN,
red (open term ktrm (tm_var (Fr x)) t) t' ->
[]; [] ++ x ~ τ1 ⊢ t' : τ2
forallt' : term LN,
red (tm_abs τ1 t) t' -> []; [] ⊢ t' : ty_ar τ1 τ2
forallt' : term LN,
red (tm_app t1 t2) t' -> []; [] ⊢ t' : τ2
L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
[] ++ x ~ tt; [] ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ H0: forallx : AtomSet.elt,
x `notin` L ->
[] ++ x ~ tt = [] ->
[] = [] ->
forallt' : term LN,
red (open term ktyp (ty_v (Fr x)) t) t' ->
[] ++ x ~ tt; [] ⊢ t'
: open typ ktyp (ty_v (Fr x)) τ
forallt' : term LN,
red (tm_tab t) t' -> []; [] ⊢ t' : ty_univ τ
t: term LN τ1, τ2: typ LN j: []; [] ⊢ t : ty_univ τ2 IHj: [] = [] ->
[] = [] ->
forallt' : term LN,
red t t' -> []; [] ⊢ t' : ty_univ τ2 H: ok_type [] τ1
forallt' : term LN,
red (tm_tap t τ1) t' ->
[]; [] ⊢ t' : open typ ktyp τ1 τ2
x: atom τ: typ LN H0: ok_type_ctx [] [] H: ok_kind_ctx [] H1: (x, τ) ∈ []
forallt' : term LN,
red (tm_var (Fr x)) t' -> []; [] ⊢ t' : τ
inversion1.
L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
[]; [] ++ x ~ τ1
⊢ open term ktrm (tm_var (Fr x)) t : τ2 H0: forallx : AtomSet.elt,
x `notin` L ->
[] = [] ->
[] ++ x ~ τ1 = [] ->
forallt' : term LN,
red (open term ktrm (tm_var (Fr x)) t) t' ->
[]; [] ++ x ~ τ1 ⊢ t' : τ2
forallt' : term LN,
red (tm_abs τ1 t) t' -> []; [] ⊢ t' : ty_ar τ1 τ2
forallt' : term LN,
red (tm_app t1 t2) t' -> []; [] ⊢ t' : τ2
t1, t2: term LN τ1, τ2: typ LN j1: []; [] ⊢ t1 : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 IHj1: [] = [] ->
[] = [] ->
forallt' : term LN,
red t1 t' -> []; [] ⊢ t' : ty_ar τ1 τ2 IHj2: [] = [] ->
[] = [] ->
forallt' : term LN,
red t2 t' -> []; [] ⊢ t' : τ1 t1': term LN H: red (tm_app t1 t2) (tm_app t1' t2) H3: red t1 t1'
[]; [] ⊢ tm_app t1' t2 : τ2
t1, t2: term LN τ1, τ2: typ LN j1: []; [] ⊢ t1 : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 IHj1: [] = [] ->
[] = [] ->
forallt' : term LN,
red t1 t' -> []; [] ⊢ t' : ty_ar τ1 τ2 IHj2: [] = [] ->
[] = [] ->
forallt' : term LN,
red t2 t' -> []; [] ⊢ t' : τ1 t2': term LN H: red (tm_app t1 t2) (tm_app t1 t2') H2: value t1 H4: red t2 t2'
[]; [] ⊢ tm_app t1 t2' : τ2
t2: term LN τ1, τ2, T: typ LN t0: term LN j1: []; [] ⊢ tm_abs T t0 : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 IHj1: [] = [] ->
[] = [] ->
forallt' : term LN,
red (tm_abs T t0) t' ->
[]; [] ⊢ t' : ty_ar τ1 τ2 IHj2: [] = [] ->
[] = [] ->
forallt' : term LN,
red t2 t' -> []; [] ⊢ t' : τ1 H: red (tm_app (tm_abs T t0) t2)
(open term ktrm t2 t0) H3: value t2
[]; [] ⊢ open term ktrm t2 t0 : τ2
t1, t2: term LN τ1, τ2: typ LN j1: []; [] ⊢ t1 : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 IHj1: [] = [] ->
[] = [] ->
forallt' : term LN,
red t1 t' -> []; [] ⊢ t' : ty_ar τ1 τ2 IHj2: [] = [] ->
[] = [] ->
forallt' : term LN,
red t2 t' -> []; [] ⊢ t' : τ1 t1': term LN H: red (tm_app t1 t2) (tm_app t1' t2) H3: red t1 t1'
[]; [] ⊢ tm_app t1' t2 : τ2
eauto using Judgment.
t1, t2: term LN τ1, τ2: typ LN j1: []; [] ⊢ t1 : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 IHj1: [] = [] ->
[] = [] ->
forallt' : term LN,
red t1 t' -> []; [] ⊢ t' : ty_ar τ1 τ2 IHj2: [] = [] ->
[] = [] ->
forallt' : term LN,
red t2 t' -> []; [] ⊢ t' : τ1 t2': term LN H: red (tm_app t1 t2) (tm_app t1 t2') H2: value t1 H4: red t2 t2'
[]; [] ⊢ tm_app t1 t2' : τ2
eauto using Judgment.
t2: term LN τ1, τ2, T: typ LN t0: term LN j1: []; [] ⊢ tm_abs T t0 : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 IHj1: [] = [] ->
[] = [] ->
forallt' : term LN,
red (tm_abs T t0) t' ->
[]; [] ⊢ t' : ty_ar τ1 τ2 IHj2: [] = [] ->
[] = [] ->
forallt' : term LN,
red t2 t' -> []; [] ⊢ t' : τ1 H: red (tm_app (tm_abs T t0) t2)
(open term ktrm t2 t0) H3: value t2
[]; [] ⊢ open term ktrm t2 t0 : τ2
t2: term LN τ1, τ2: typ LN t0: term LN j2: []; [] ⊢ t2 : τ1 IHj1: [] = [] ->
[] = [] ->
forallt' : term LN,
red (tm_abs τ1 t0) t' ->
[]; [] ⊢ t' : ty_ar τ1 τ2 IHj2: [] = [] ->
[] = [] ->
forallt' : term LN,
red t2 t' -> []; [] ⊢ t' : τ1 H: red (tm_app (tm_abs τ1 t0) t2)
(open term ktrm t2 t0) H3: value t2 L: AtomSet.t H4: forallx : AtomSet.elt,
x `notin` L ->
[]; [] ++ x ~ τ1
⊢ open term ktrm (tm_var (Fr x)) t0 : τ2
[]; [] ⊢ open term ktrm t2 t0 : τ2
t2: term LN τ1, τ2: typ LN t0: term LN j2: []; [] ⊢ t2 : τ1 IHj1: [] = [] ->
[] = [] ->
forallt' : term LN,
red (tm_abs τ1 t0) t' ->
[]; [] ⊢ t' : ty_ar τ1 τ2 IHj2: [] = [] ->
[] = [] ->
forallt' : term LN,
red t2 t' -> []; [] ⊢ t' : τ1 H: red (tm_app (tm_abs τ1 t0) t2)
(open term ktrm t2 t0) H3: value t2 L: AtomSet.t hyp: forallx : AtomSet.elt,
x `notin` L ->
[]; [] ++ x ~ τ1
⊢ open term ktrm (tm_var (Fr x)) t0 : τ2
[]; [] ⊢ open term ktrm t2 t0 : τ2
t2: term LN τ1, τ2: typ LN t0: term LN j2: []; [] ⊢ t2 : τ1 IHj1: [] = [] ->
[] = [] ->
forallt' : term LN,
red (tm_abs τ1 t0) t' ->
[]; [] ⊢ t' : ty_ar τ1 τ2 IHj2: [] = [] ->
[] = [] ->
forallt' : term LN,
red t2 t' -> []; [] ⊢ t' : τ1 H: red (tm_app (tm_abs τ1 t0) t2)
(open term ktrm t2 t0) H3: value t2 L: AtomSet.t hyp: forallx : AtomSet.elt,
x `notin` L ->
[]; [] ++ x ~ τ1
⊢ open term ktrm (tm_var (Fr x)) t0 : τ2 e: atom Hfresh: e `notin` (L ∪ FV term ktrm t0)
[]; [] ⊢ open term ktrm t2 t0 : τ2
t2: term LN τ1, τ2: typ LN t0: term LN j2: []; [] ⊢ t2 : τ1 IHj1: [] = [] ->
[] = [] ->
forallt' : term LN,
red (tm_abs τ1 t0) t' ->
[]; [] ⊢ t' : ty_ar τ1 τ2 IHj2: [] = [] ->
[] = [] ->
forallt' : term LN,
red t2 t' -> []; [] ⊢ t' : τ1 H: red (tm_app (tm_abs τ1 t0) t2)
(open term ktrm t2 t0) H3: value t2 L: AtomSet.t hyp: forallx : AtomSet.elt,
x `notin` L ->
[]; [] ++ x ~ τ1
⊢ open term ktrm (tm_var (Fr x)) t0 : τ2 e: atom Hfresh: e `notin` (L ∪ FV term ktrm t0)
[]; []
⊢ subst term ktrm e t2
(open term ktrm (mret SystemF ktrm (Fr e)) t0)
: τ2
t2: term LN τ1, τ2: typ LN t0: term LN j2: []; [] ⊢ t2 : τ1 IHj1: [] = [] ->
[] = [] ->
forallt' : term LN,
red (tm_abs τ1 t0) t' ->
[]; [] ⊢ t' : ty_ar τ1 τ2 IHj2: [] = [] ->
[] = [] ->
forallt' : term LN,
red t2 t' -> []; [] ⊢ t' : τ1 H: red (tm_app (tm_abs τ1 t0) t2)
(open term ktrm t2 t0) H3: value t2 L: AtomSet.t hyp: forallx : AtomSet.elt,
x `notin` L ->
[]; [] ++ x ~ τ1
⊢ open term ktrm (tm_var (Fr x)) t0 : τ2 e: atom Hfresh: e `notin` (L ∪ FV term ktrm t0)
[]; [] ++ e ~ τ1
⊢ open term ktrm (mret SystemF ktrm (Fr e)) t0 : τ2
t2: term LN τ1, τ2: typ LN t0: term LN j2: []; [] ⊢ t2 : τ1 IHj1: [] = [] ->
[] = [] ->
forallt' : term LN,
red (tm_abs τ1 t0) t' ->
[]; [] ⊢ t' : ty_ar τ1 τ2 IHj2: [] = [] ->
[] = [] ->
forallt' : term LN,
red t2 t' -> []; [] ⊢ t' : τ1 H: red (tm_app (tm_abs τ1 t0) t2)
(open term ktrm t2 t0) H3: value t2 L: AtomSet.t hyp: forallx : AtomSet.elt,
x `notin` L ->
[]; [] ++ x ~ τ1
⊢ open term ktrm (tm_var (Fr x)) t0 : τ2 e: atom Hfresh: e `notin` (L ∪ FV term ktrm t0)
e `notin` L
fsetdec.
L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
[] ++ x ~ tt; [] ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ H0: forallx : AtomSet.elt,
x `notin` L ->
[] ++ x ~ tt = [] ->
[] = [] ->
forallt' : term LN,
red (open term ktyp (ty_v (Fr x)) t) t' ->
[] ++ x ~ tt; [] ⊢ t'
: open typ ktyp (ty_v (Fr x)) τ
forallt' : term LN,
red (tm_tab t) t' -> []; [] ⊢ t' : ty_univ τ
inversion1.
t: term LN τ1, τ2: typ LN j: []; [] ⊢ t : ty_univ τ2 IHj: [] = [] ->
[] = [] ->
forallt' : term LN,
red t t' -> []; [] ⊢ t' : ty_univ τ2 H: ok_type [] τ1
forallt' : term LN,
red (tm_tap t τ1) t' ->
[]; [] ⊢ t' : open typ ktyp τ1 τ2
t: term LN τ1, τ2: typ LN j: []; [] ⊢ t : ty_univ τ2 IHj: [] = [] ->
[] = [] ->
forallt' : term LN,
red t t' -> []; [] ⊢ t' : ty_univ τ2 H: ok_type [] τ1 t'0: term LN H0: red (tm_tap t τ1) (tm_tap t'0 τ1) H4: red t t'0
[]; [] ⊢ tm_tap t'0 τ1 : open typ ktyp τ1 τ2
τ1, τ2: typ LN t0: term LN IHj: [] = [] ->
[] = [] ->
forallt' : term LN,
red (tm_tab t0) t' -> []; [] ⊢ t' : ty_univ τ2 j: []; [] ⊢ tm_tab t0 : ty_univ τ2 H: ok_type [] τ1 H0: red (tm_tap (tm_tab t0) τ1)
(open term ktyp τ1 t0)
[]; [] ⊢ open term ktyp τ1 t0 : open typ ktyp τ1 τ2
t: term LN τ1, τ2: typ LN j: []; [] ⊢ t : ty_univ τ2 IHj: [] = [] ->
[] = [] ->
forallt' : term LN,
red t t' -> []; [] ⊢ t' : ty_univ τ2 H: ok_type [] τ1 t'0: term LN H0: red (tm_tap t τ1) (tm_tap t'0 τ1) H4: red t t'0
[]; [] ⊢ tm_tap t'0 τ1 : open typ ktyp τ1 τ2
eauto using Judgment.
τ1, τ2: typ LN t0: term LN IHj: [] = [] ->
[] = [] ->
forallt' : term LN,
red (tm_tab t0) t' -> []; [] ⊢ t' : ty_univ τ2 j: []; [] ⊢ tm_tab t0 : ty_univ τ2 H: ok_type [] τ1 H0: red (tm_tap (tm_tab t0) τ1)
(open term ktyp τ1 t0)
[]; [] ⊢ open term ktyp τ1 t0 : open typ ktyp τ1 τ2
τ1, τ2: typ LN t0: term LN IHj: [] = [] ->
[] = [] ->
forallt' : term LN,
red (tm_tab t0) t' -> []; [] ⊢ t' : ty_univ τ2 H: ok_type [] τ1 H0: red (tm_tap (tm_tab t0) τ1)
(open term ktyp τ1 t0) L: AtomSet.t H5: forallx : AtomSet.elt,
x `notin` L ->
[] ++ x ~ tt; []
⊢ open term ktyp (ty_v (Fr x)) t0
: open typ ktyp (ty_v (Fr x)) τ2
[]; [] ⊢ open term ktyp τ1 t0 : open typ ktyp τ1 τ2
τ1, τ2: typ LN t0: term LN IHj: [] = [] ->
[] = [] ->
forallt' : term LN,
red (tm_tab t0) t' -> []; [] ⊢ t' : ty_univ τ2 H: ok_type [] τ1 H0: red (tm_tap (tm_tab t0) τ1)
(open term ktyp τ1 t0) L: AtomSet.t H5: forallx : AtomSet.elt,
x `notin` L ->
[] ++ x ~ tt; []
⊢ open term ktyp (ty_v (Fr x)) t0
: open typ ktyp (ty_v (Fr x)) τ2 e: atom Hfresh: e
`notin` (L
∪ (FV term ktyp t0 ∪ FV typ ktyp τ2))
[]; [] ⊢ open term ktyp τ1 t0 : open typ ktyp τ1 τ2
τ1, τ2: typ LN t0: term LN IHj: [] = [] ->
[] = [] ->
forallt' : term LN,
red (tm_tab t0) t' -> []; [] ⊢ t' : ty_univ τ2 H: ok_type [] τ1 H0: red (tm_tap (tm_tab t0) τ1)
(open term ktyp τ1 t0) L: AtomSet.t H5: forallx : AtomSet.elt,
x `notin` L ->
[] ++ x ~ tt; []
⊢ open term ktyp (ty_v (Fr x)) t0
: open typ ktyp (ty_v (Fr x)) τ2 e: atom Hfresh: e
`notin` (L
∪ (FV term ktyp t0 ∪ FV typ ktyp τ2))
[]; []
⊢ subst term ktyp e τ1
(open term ktyp (mret SystemF ktyp (Fr e)) t0)
: open typ ktyp τ1 τ2
τ1, τ2: typ LN t0: term LN IHj: [] = [] ->
[] = [] ->
forallt' : term LN,
red (tm_tab t0) t' -> []; [] ⊢ t' : ty_univ τ2 H: ok_type [] τ1 H0: red (tm_tap (tm_tab t0) τ1)
(open term ktyp τ1 t0) L: AtomSet.t H5: forallx : AtomSet.elt,
x `notin` L ->
[] ++ x ~ tt; []
⊢ open term ktyp (ty_v (Fr x)) t0
: open typ ktyp (ty_v (Fr x)) τ2 e: atom Hfresh: e
`notin` (L
∪ (FV term ktyp t0 ∪ FV typ ktyp τ2))
[]; []
⊢ subst term ktyp e τ1
(open term ktyp (mret SystemF ktyp (Fr e)) t0)
: subst typ ktyp e τ1
(open typ ktyp (mret SystemF ktyp (Fr e)) τ2)
τ1, τ2: typ LN t0: term LN IHj: [] = [] ->
[] = [] ->
forallt' : term LN,
red (tm_tab t0) t' -> []; [] ⊢ t' : ty_univ τ2 H: ok_type [] τ1 H0: red (tm_tap (tm_tab t0) τ1)
(open term ktyp τ1 t0) L: AtomSet.t H5: forallx : AtomSet.elt,
x `notin` L ->
[] ++ x ~ tt; []
⊢ open term ktyp (ty_v (Fr x)) t0
: open typ ktyp (ty_v (Fr x)) τ2 e: atom Hfresh: e
`notin` (L
∪ (FV term ktyp t0 ∪ FV typ ktyp τ2))
[]; envmap (subst typ ktyp e τ1) []
⊢ subst term ktyp e τ1
(open term ktyp (mret SystemF ktyp (Fr e)) t0)
: subst typ ktyp e τ1
(open typ ktyp (mret SystemF ktyp (Fr e)) τ2)
τ1, τ2: typ LN t0: term LN IHj: [] = [] ->
[] = [] ->
forallt' : term LN,
red (tm_tab t0) t' -> []; [] ⊢ t' : ty_univ τ2 H: ok_type [] τ1 H0: red (tm_tap (tm_tab t0) τ1)
(open term ktyp τ1 t0) L: AtomSet.t H5: forallx : AtomSet.elt,
x `notin` L ->
[] ++ x ~ tt; []
⊢ open term ktyp (ty_v (Fr x)) t0
: open typ ktyp (ty_v (Fr x)) τ2 e: atom Hfresh: e
`notin` (L
∪ (FV term ktyp t0 ∪ FV typ ktyp τ2))
ok_type [] τ1
τ1, τ2: typ LN t0: term LN IHj: [] = [] ->
[] = [] ->
forallt' : term LN,
red (tm_tab t0) t' -> []; [] ⊢ t' : ty_univ τ2 H: ok_type [] τ1 H0: red (tm_tap (tm_tab t0) τ1)
(open term ktyp τ1 t0) L: AtomSet.t H5: forallx : AtomSet.elt,
x `notin` L ->
[] ++ x ~ tt; []
⊢ open term ktyp (ty_v (Fr x)) t0
: open typ ktyp (ty_v (Fr x)) τ2 e: atom Hfresh: e
`notin` (L
∪ (FV term ktyp t0 ∪ FV typ ktyp τ2))
[] ++ e ~ tt; []
⊢ open term ktyp (mret SystemF ktyp (Fr e)) t0
: open typ ktyp (mret SystemF ktyp (Fr e)) τ2
τ1, τ2: typ LN t0: term LN IHj: [] = [] ->
[] = [] ->
forallt' : term LN,
red (tm_tab t0) t' -> []; [] ⊢ t' : ty_univ τ2 H: ok_type [] τ1 H0: red (tm_tap (tm_tab t0) τ1)
(open term ktyp τ1 t0) L: AtomSet.t H5: forallx : AtomSet.elt,
x `notin` L ->
[] ++ x ~ tt; []
⊢ open term ktyp (ty_v (Fr x)) t0
: open typ ktyp (ty_v (Fr x)) τ2 e: atom Hfresh: e
`notin` (L
∪ (FV term ktyp t0 ∪ FV typ ktyp τ2))
ok_type [] τ1
assumption.
τ1, τ2: typ LN t0: term LN IHj: [] = [] ->
[] = [] ->
forallt' : term LN,
red (tm_tab t0) t' -> []; [] ⊢ t' : ty_univ τ2 H: ok_type [] τ1 H0: red (tm_tap (tm_tab t0) τ1)
(open term ktyp τ1 t0) L: AtomSet.t H5: forallx : AtomSet.elt,
x `notin` L ->
[] ++ x ~ tt; []
⊢ open term ktyp (ty_v (Fr x)) t0
: open typ ktyp (ty_v (Fr x)) τ2 e: atom Hfresh: e
`notin` (L
∪ (FV term ktyp t0 ∪ FV typ ktyp τ2))
[] ++ e ~ tt; []
⊢ open term ktyp (mret SystemF ktyp (Fr e)) t0
: open typ ktyp (mret SystemF ktyp (Fr e)) τ2
τ1, τ2: typ LN t0: term LN IHj: [] = [] ->
[] = [] ->
forallt' : term LN,
red (tm_tab t0) t' -> []; [] ⊢ t' : ty_univ τ2 H: ok_type [] τ1 H0: red (tm_tap (tm_tab t0) τ1)
(open term ktyp τ1 t0) L: AtomSet.t H5: forallx : AtomSet.elt,
x `notin` L ->
[] ++ x ~ tt; []
⊢ open term ktyp (ty_v (Fr x)) t0
: open typ ktyp (ty_v (Fr x)) τ2 e: atom Hfresh: e
`notin` (L
∪ (FV term ktyp t0 ∪ FV typ ktyp τ2))
[] ++ e ~ tt; []
⊢ open term ktyp (mret SystemF ktyp (Fr e)) t0
: open typ ktyp (mret SystemF ktyp (Fr e)) τ2
τ1, τ2: typ LN t0: term LN IHj: [] = [] ->
[] = [] ->
forallt' : term LN,
red (tm_tab t0) t' -> []; [] ⊢ t' : ty_univ τ2 H: ok_type [] τ1 H0: red (tm_tap (tm_tab t0) τ1)
(open term ktyp τ1 t0) L: AtomSet.t H5: forallx : AtomSet.elt,
x `notin` L ->
[] ++ x ~ tt; []
⊢ open term ktyp (ty_v (Fr x)) t0
: open typ ktyp (ty_v (Fr x)) τ2 e: atom Hfresh: e
`notin` (L
∪ (FV term ktyp t0 ∪ FV typ ktyp τ2))
e `notin` L
fsetdec.}Qed.
progress
progress
t: term LN τ: typ LN j: []; [] ⊢ t : τ
value t \/ (existst' : term LN, red t t')
t: term LN τ: typ LN Delta: list (atom * unit) HeqDelta: Delta = [] j: Delta; [] ⊢ t : τ
value t \/ (existst' : term LN, red t t')
t: term LN τ: typ LN Delta: list (atom * unit) HeqDelta: Delta = [] Gamma: list (atom * typ LN) HeqGamma: Gamma = [] j: Delta; Gamma ⊢ t : τ
value t \/ (existst' : term LN, red t t')
x: atom τ: typ LN H0: ok_type_ctx [] [] H: ok_kind_ctx [] H1: (x, τ) ∈ []
value (tm_var (Fr x)) \/
(existst' : term LN, red (tm_var (Fr x)) t')
L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
[]; [] ++ x ~ τ1
⊢ open term ktrm (tm_var (Fr x)) t : τ2 H0: forallx : AtomSet.elt,
x `notin` L ->
[] = [] ->
[] ++ x ~ τ1 = [] ->
value (open term ktrm (tm_var (Fr x)) t) \/
(existst' : term LN,
red (open term ktrm (tm_var (Fr x)) t) t')
value (tm_abs τ1 t) \/
(existst' : term LN, red (tm_abs τ1 t) t')
t1, t2: term LN τ1, τ2: typ LN j1: []; [] ⊢ t1 : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 IHj1: [] = [] ->
[] = [] ->
value t1 \/ (existst' : term LN, red t1 t') IHj2: [] = [] ->
[] = [] ->
value t2 \/ (existst' : term LN, red t2 t')
value (tm_app t1 t2) \/
(existst' : term LN, red (tm_app t1 t2) t')
L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
[] ++ x ~ tt; [] ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ H0: forallx : AtomSet.elt,
x `notin` L ->
[] ++ x ~ tt = [] ->
[] = [] ->
value (open term ktyp (ty_v (Fr x)) t) \/
(existst' : term LN,
red (open term ktyp (ty_v (Fr x)) t) t')
value (tm_tab t) \/
(existst' : term LN, red (tm_tab t) t')
t: term LN τ1, τ2: typ LN j: []; [] ⊢ t : ty_univ τ2 IHj: [] = [] ->
[] = [] ->
value t \/ (existst' : term LN, red t t') H: ok_type [] τ1
value (tm_tap t τ1) \/
(existst' : term LN, red (tm_tap t τ1) t')
x: atom τ: typ LN H0: ok_type_ctx [] [] H: ok_kind_ctx [] H1: (x, τ) ∈ []
value (tm_var (Fr x)) \/
(existst' : term LN, red (tm_var (Fr x)) t')
inversion H1.
L: AtomSet.t t: term LN τ1, τ2: typ LN H: forallx : AtomSet.elt,
x `notin` L ->
[]; [] ++ x ~ τ1
⊢ open term ktrm (tm_var (Fr x)) t : τ2 H0: forallx : AtomSet.elt,
x `notin` L ->
[] = [] ->
[] ++ x ~ τ1 = [] ->
value (open term ktrm (tm_var (Fr x)) t) \/
(existst' : term LN,
red (open term ktrm (tm_var (Fr x)) t) t')
value (tm_abs τ1 t) \/
(existst' : term LN, red (tm_abs τ1 t) t')
left; auto using value.
t1, t2: term LN τ1, τ2: typ LN j1: []; [] ⊢ t1 : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 IHj1: [] = [] ->
[] = [] ->
value t1 \/ (existst' : term LN, red t1 t') IHj2: [] = [] ->
[] = [] ->
value t2 \/ (existst' : term LN, red t2 t')
value (tm_app t1 t2) \/
(existst' : term LN, red (tm_app t1 t2) t')
t1, t2: term LN τ1, τ2: typ LN j1: []; [] ⊢ t1 : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 IHj1: [] = [] ->
[] = [] ->
value t1 \/ (existst' : term LN, red t1 t') IHj2: [] = [] ->
[] = [] ->
value t2 \/ (existst' : term LN, red t2 t')
existst' : term LN, red (tm_app t1 t2) t'
t1, t2: term LN τ1, τ2: typ LN j1: []; [] ⊢ t1 : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 IHj1: value t1 \/ (existst' : term LN, red t1 t') IHj2: [] = [] ->
[] = [] ->
value t2 \/ (existst' : term LN, red t2 t')
existst' : term LN, red (tm_app t1 t2) t'
t1, t2: term LN τ1, τ2: typ LN j1: []; [] ⊢ t1 : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 IHj1: value t1 \/ (existst' : term LN, red t1 t') IHj2: value t2 \/ (existst' : term LN, red t2 t')
existst' : term LN, red (tm_app t1 t2) t'
t1, t2: term LN τ1, τ2: typ LN j1: []; [] ⊢ t1 : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 H: value t1 H0: value t2
existst' : term LN, red (tm_app t1 t2) t'
t1, t2: term LN τ1, τ2: typ LN j1: []; [] ⊢ t1 : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 H: value t1 H0: existst' : term LN, red t2 t'
existst' : term LN, red (tm_app t1 t2) t'
t1, t2: term LN τ1, τ2: typ LN j1: []; [] ⊢ t1 : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 H: existst' : term LN, red t1 t' H0: value t2
existst' : term LN, red (tm_app t1 t2) t'
t1, t2: term LN τ1, τ2: typ LN j1: []; [] ⊢ t1 : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 H: existst' : term LN, red t1 t' H0: existst' : term LN, red t2 t'
existst' : term LN, red (tm_app t1 t2) t'
t1, t2: term LN τ1, τ2: typ LN j1: []; [] ⊢ t1 : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 H: value t1 H0: value t2
existst' : term LN, red (tm_app t1 t2) t'
t2: term LN τ1, τ2, T: typ LN t: term LN j1: []; [] ⊢ tm_abs T t : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 H0: value t2
existst' : term LN, red (tm_app (tm_abs T t) t2) t'
t2: term LN τ1, τ2: typ LN t: term LN j1: []; [] ⊢ tm_tab t : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 H0: value t2
existst' : term LN, red (tm_app (tm_tab t) t2) t'
t2: term LN τ1, τ2, T: typ LN t: term LN j1: []; [] ⊢ tm_abs T t : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 H0: value t2
existst' : term LN, red (tm_app (tm_abs T t) t2) t'
eauto usingred.
t2: term LN τ1, τ2: typ LN t: term LN j1: []; [] ⊢ tm_tab t : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 H0: value t2
existst' : term LN, red (tm_app (tm_tab t) t2) t'
t2: term LN τ1, τ2: typ LN t: term LN j1: []; [] ⊢ tm_tab t : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 H0: value t2
existst' : term LN, red (tm_app (tm_tab t) t2) t'
t2: term LN τ1, τ2: typ LN t: term LN j1: []; [] ⊢ tm_tab t : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 H0: value t2
False
inverts j1.}
t1, t2: term LN τ1, τ2: typ LN j1: []; [] ⊢ t1 : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 H: value t1 H0: existst' : term LN, red t2 t'
existst' : term LN, red (tm_app t1 t2) t'
t2: term LN τ1, τ2, T: typ LN t: term LN j1: []; [] ⊢ tm_abs T t : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 H0: existst' : term LN, red t2 t'
existst' : term LN, red (tm_app (tm_abs T t) t2) t'
t2: term LN τ1, τ2: typ LN t: term LN j1: []; [] ⊢ tm_tab t : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 H0: existst' : term LN, red t2 t'
existst' : term LN, red (tm_app (tm_tab t) t2) t'
t2: term LN τ1, τ2, T: typ LN t: term LN j1: []; [] ⊢ tm_abs T t : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 H0: existst' : term LN, red t2 t'
existst' : term LN, red (tm_app (tm_abs T t) t2) t'
t2: term LN τ1, τ2, T: typ LN t: term LN j1: []; [] ⊢ tm_abs T t : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 x: term LN H: red t2 x
existst' : term LN, red (tm_app (tm_abs T t) t2) t'
eauto usingred, value.
t2: term LN τ1, τ2: typ LN t: term LN j1: []; [] ⊢ tm_tab t : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 H0: existst' : term LN, red t2 t'
existst' : term LN, red (tm_app (tm_tab t) t2) t'
t2: term LN τ1, τ2: typ LN t: term LN j1: []; [] ⊢ tm_tab t : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 H0: existst' : term LN, red t2 t'
existst' : term LN, red (tm_app (tm_tab t) t2) t'
inverts j1.}
t1, t2: term LN τ1, τ2: typ LN j1: []; [] ⊢ t1 : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 H: existst' : term LN, red t1 t' H0: value t2
existst' : term LN, red (tm_app t1 t2) t'
t1, t2: term LN τ1, τ2: typ LN j1: []; [] ⊢ t1 : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 x: term LN H1: red t1 x H0: value t2
existst' : term LN, red (tm_app t1 t2) t'
eauto usingred.
t1, t2: term LN τ1, τ2: typ LN j1: []; [] ⊢ t1 : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 H: existst' : term LN, red t1 t' H0: existst' : term LN, red t2 t'
existst' : term LN, red (tm_app t1 t2) t'
t1, t2: term LN τ1, τ2: typ LN j1: []; [] ⊢ t1 : ty_ar τ1 τ2 j2: []; [] ⊢ t2 : τ1 x0: term LN H0: red t1 x0 x: term LN H1: red t2 x
existst' : term LN, red (tm_app t1 t2) t'
eauto usingred.
L: AtomSet.t τ: typ LN t: term LN H: forallx : AtomSet.elt,
x `notin` L ->
[] ++ x ~ tt; [] ⊢ open term ktyp (ty_v (Fr x)) t
: open typ ktyp (ty_v (Fr x)) τ H0: forallx : AtomSet.elt,
x `notin` L ->
[] ++ x ~ tt = [] ->
[] = [] ->
value (open term ktyp (ty_v (Fr x)) t) \/
(existst' : term LN,
red (open term ktyp (ty_v (Fr x)) t) t')
value (tm_tab t) \/
(existst' : term LN, red (tm_tab t) t')
left; auto using value.
t: term LN τ1, τ2: typ LN j: []; [] ⊢ t : ty_univ τ2 IHj: [] = [] ->
[] = [] ->
value t \/ (existst' : term LN, red t t') H: ok_type [] τ1
value (tm_tap t τ1) \/
(existst' : term LN, red (tm_tap t τ1) t')
t: term LN τ1, τ2: typ LN j: []; [] ⊢ t : ty_univ τ2 IHj: [] = [] ->
[] = [] ->
value t \/ (existst' : term LN, red t t') H: ok_type [] τ1
existst' : term LN, red (tm_tap t τ1) t'
t: term LN τ1, τ2: typ LN j: []; [] ⊢ t : ty_univ τ2 IHj: value t \/ (existst' : term LN, red t t') H: ok_type [] τ1
existst' : term LN, red (tm_tap t τ1) t'
t: term LN τ1, τ2: typ LN j: []; [] ⊢ t : ty_univ τ2 H: ok_type [] τ1 H0: value t
existst' : term LN, red (tm_tap t τ1) t'
t: term LN τ1, τ2: typ LN j: []; [] ⊢ t : ty_univ τ2 H: ok_type [] τ1 H0: existst' : term LN, red t t'
existst' : term LN, red (tm_tap t τ1) t'
t: term LN τ1, τ2: typ LN j: []; [] ⊢ t : ty_univ τ2 H: ok_type [] τ1 H0: value t
existst' : term LN, red (tm_tap t τ1) t'
τ1, τ2, T: typ LN t0: term LN j: []; [] ⊢ tm_abs T t0 : ty_univ τ2 H: ok_type [] τ1
existst' : term LN, red (tm_tap (tm_abs T t0) τ1) t'
τ1, τ2: typ LN t0: term LN j: []; [] ⊢ tm_tab t0 : ty_univ τ2 H: ok_type [] τ1
existst' : term LN, red (tm_tap (tm_tab t0) τ1) t'
τ1, τ2, T: typ LN t0: term LN j: []; [] ⊢ tm_abs T t0 : ty_univ τ2 H: ok_type [] τ1
existst' : term LN, red (tm_tap (tm_abs T t0) τ1) t'
inverts j.
τ1, τ2: typ LN t0: term LN j: []; [] ⊢ tm_tab t0 : ty_univ τ2 H: ok_type [] τ1
existst' : term LN, red (tm_tap (tm_tab t0) τ1) t'
τ1, τ2: typ LN t0: term LN j: []; [] ⊢ tm_tab t0 : ty_univ τ2 H: ok_type [] τ1
existst' : term LN, red (tm_tap (tm_tab t0) τ1) t'
eauto usingred.}
t: term LN τ1, τ2: typ LN j: []; [] ⊢ t : ty_univ τ2 H: ok_type [] τ1 H0: existst' : term LN, red t t'
existst' : term LN, red (tm_tap t τ1) t'
t: term LN τ1, τ2: typ LN j: []; [] ⊢ t : ty_univ τ2 H: ok_type [] τ1 x: term LN H1: red t x
existst' : term LN, red (tm_tap t τ1) t'
eauto usingred.Qed.
Axioms:
base_typ : Type
(* Axioms:base_typ : Type*)
Axioms:
propositional_extensionality
: forallPQ : Prop, P <-> Q -> P = Q
functional_extensionality_dep
: forall (A : Type) (B : A -> Type)
(fg : forallx : A, B x),
(forallx : A, f x = g x) -> f = g
Eqdep.Eq_rect_eq.eq_rect_eq
: forall (U : Type) (p : U) (Q : U -> Type)
(x : Q p) (h : p = p), x = rew [Q] h in x
base_typ : Type
(*Axioms:propositional_extensionality : forall P Q : Prop, P <-> Q -> P = Qfunctional_extensionality_dep : forall (A : Type) (B : A -> Type) (f g : forall x : A, B x), (forall x : A, f x = g x) -> f = gEqdep.Eq_rect_eq.eq_rect_eq : forall (U : Type) (p : U) (Q : U -> Type) (x : Q p) (h : p = p), x = rew [Q] h in xbase_typ : Type*)