(* $Id: Fsub.thy,v 1.35 2007/07/11 09:36:07 berghofe Exp $ *) (*<*) theory Fsub imports "../Nominal" begin (*>*) text{* Authors: Christian Urban, Benjamin Pierce, Dimitrios Vytiniotis Stephanie Weirich and Steve Zdancewic with great help from Stefan Berghofer and Markus Wenzel. *} section {* Types for Names, Nominal Datatype Declaration for Types and Terms *} text {* The main point of this solution is to use names everywhere (be they bound, binding or free). In System \FSUB{} there are two kinds of names corresponding to type-variables and to term-variables. These two kinds of names are represented in the nominal datatype package as atom-types @{text "tyvrs"} and @{text "vrs"}: *} atom_decl tyvrs vrs text{* There are numerous facts that come with this declaration: for example that there are infinitely many elements in @{text "tyvrs"} and @{text "vrs"}. *} text{* The constructors for types and terms in System \FSUB{} contain abstractions over type-variables and term-variables. The nominal datatype-package uses @{text "«…»…"} to indicate where abstractions occur. *} nominal_datatype ty = Tvar "tyvrs" | Top | Arrow "ty" "ty" ("_ -> _" [100,100] 100) | Forall "«tyvrs»ty" "ty" nominal_datatype trm = Var "vrs" | Lam "«vrs»trm" "ty" | Tabs "«tyvrs»trm" "ty" | App "trm" "trm" | Tapp "trm" "ty" text {* To be polite to the eye, some more familiar notation is introduced. Because of the change in the order of arguments, one needs to use translation rules, instead of syntax annotations at the term-constructors as given above for @{term "Arrow"}. *} syntax Forall_syn :: "tyvrs => ty => ty => ty" ("∀[_<:_]._" [100,100,100] 100) Lam_syn :: "vrs => ty => trm => trm" ("Lam [_:_]._" [100,100,100] 100) Tabs_syn :: "tyvrs => ty => trm => trm" ("Tabs [_<:_]._" [100,100,100] 100) translations "∀[X<:T1].T2" \<rightleftharpoons> "ty.Forall X T2 T1" "Lam [x:T].t" \<rightleftharpoons> "trm.Lam x t T" "Tabs [X<:T].t" \<rightleftharpoons> "trm.Tabs X t T" text {* Again there are numerous facts that are proved automatically for @{typ "ty"} and @{typ "trm"}: for example that the set of free variables, i.e.~the @{text "support"}, is finite. However note that nominal-datatype declarations do \emph{not} define ``classical" constructor-based datatypes, but rather define $\alpha$-equivalence classes---we can for example show that $\alpha$-equivalent @{typ "ty"}s and @{typ "trm"}s are equal: *} lemma alpha_illustration: shows "∀[X<:T].(Tvar X) = ∀[Y<:T].(Tvar Y)" and "Lam [x:T].(Var x) = Lam [y:T].(Var y)" by (simp_all add: ty.inject trm.inject alpha calc_atm fresh_atm) section {* SubTyping Contexts *} types ty_context = "(tyvrs×ty) list" text {* Typing contexts are represented as lists that ``grow" on the left; we thereby deviating from the convention in the POPLmark-paper. The lists contain pairs of type-variables and types (this is sufficient for Part 1A). *} text {* In order to state validity-conditions for typing-contexts, the notion of a @{text "domain"} of a typing-context is handy. *} consts "domain" :: "ty_context => tyvrs set" primrec "domain [] = {}" "domain (X#Γ) = {fst X}∪(domain Γ)" lemma domain_eqvt[eqvt]: fixes pi::"tyvrs prm" and pi'::"vrs prm" shows "pi•(domain Γ) = domain (pi•Γ)" and "pi'•(domain Γ) = domain (pi'•Γ)" by (induct Γ) (simp_all add: eqvts) lemma finite_domain: shows "finite (domain Γ)" by (induct Γ, auto) lemma domain_supp: shows "(supp (domain Γ)) = (domain Γ)" by (simp only: at_fin_set_supp at_tyvrs_inst finite_domain) lemma domain_inclusion: assumes a: "(X,T)∈set Γ" shows "X∈(domain Γ)" using a by (induct Γ, auto) lemma domain_existence: assumes a: "X∈(domain Γ)" shows "∃T.(X,T)∈set Γ" using a by (induct Γ, auto) lemma domain_append: shows "domain (Γ@Δ) = ((domain Γ) ∪ (domain Δ))" by (induct Γ, auto) lemma fresh_domain_cons: fixes X::"tyvrs" shows "X\<sharp>(domain (Y#Γ)) = (X\<sharp>(fst Y) ∧ X\<sharp>(domain Γ))" by (simp add: fresh_fin_insert pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst finite_domain) lemma fresh_domain: fixes X::"tyvrs" assumes a: "X\<sharp>Γ" shows "X\<sharp>(domain Γ)" using a apply(induct Γ) apply(simp add: fresh_set_empty) apply(simp only: fresh_domain_cons) apply(auto simp add: fresh_prod fresh_list_cons) done text {* Not all lists of type @{typ "ty_context"} are well-formed. One condition requires that in @{term "(X,S)#Γ"} all free variables of @{term "S"} must be in the @{term "domain"} of @{term "Γ"}, that is @{term "S"} must be @{text "closed"} in @{term "Γ"}. The set of free variables of @{term "S"} is the @{text "support"} of @{term "S"}. *} constdefs "closed_in" :: "ty => ty_context => bool" ("_ closed'_in _" [100,100] 100) "S closed_in Γ ≡ (supp S)⊆(domain Γ)" lemma closed_in_eqvt[eqvt]: fixes pi::"tyvrs prm" assumes a: "S closed_in Γ" shows "(pi•S) closed_in (pi•Γ)" using a proof (unfold "closed_in_def") assume "supp S ⊆ (domain Γ)" hence "pi•(supp S) ⊆ pi•(domain Γ)" by (simp add: pt_subseteq_eqvt[OF pt_tyvrs_inst, OF at_tyvrs_inst]) thus "(supp (pi•S)) ⊆ (domain (pi•Γ))" by (simp add: domain_eqvt pt_perm_supp[OF pt_tyvrs_inst, OF at_tyvrs_inst]) qed lemma ty_vrs_prm_simp: fixes pi::"vrs prm" and S::"ty" shows "pi•S = S" by (induct S rule: ty.weak_induct) (auto simp add: calc_atm) lemma ty_context_vrs_prm_simp: fixes pi::"vrs prm" and Γ::"ty_context" shows "pi•Γ = Γ" by (induct Γ) (auto simp add: calc_atm ty_vrs_prm_simp) lemma closed_in_eqvt'[eqvt]: fixes pi::"vrs prm" assumes a: "S closed_in Γ" shows "(pi•S) closed_in (pi•Γ)" using a by (simp add: ty_vrs_prm_simp ty_context_vrs_prm_simp) text {* Now validity of a context is a straightforward inductive definition. *} inductive valid_rel :: "ty_context => bool" ("\<turnstile> _ ok" [100] 100) where valid_nil[simp]: "\<turnstile> [] ok" | valid_cons[simp]: "[|\<turnstile> Γ ok; X\<sharp>(domain Γ); T closed_in Γ|] ==> \<turnstile> ((X,T)#Γ) ok" equivariance valid_rel lemma validE: assumes a: "\<turnstile> ((X,T)#Γ) ok" shows "\<turnstile> Γ ok ∧ X\<sharp>(domain Γ) ∧ T closed_in Γ" using a by (cases, auto) lemma validE_append: assumes a: "\<turnstile> (Δ@Γ) ok" shows "\<turnstile> Γ ok" using a by (induct Δ, auto dest: validE) lemma replace_type: assumes a: "\<turnstile> (Δ@(X,T)#Γ) ok" and b: "S closed_in Γ" shows "\<turnstile> (Δ@(X,S)#Γ) ok" using a b apply(induct Δ) apply(auto dest!: validE intro!: valid_cons simp add: domain_append closed_in_def) done text {* Well-formed contexts have a unique type-binding for a type-variable. *} lemma uniqueness_of_ctxt: fixes Γ::"ty_context" assumes a: "\<turnstile> Γ ok" and b: "(X,T)∈set Γ" and c: "(X,S)∈set Γ" shows "T=S" using a b c proof (induct) case valid_nil thus "T=S" by simp next case valid_cons moreover { fix Γ::"ty_context" assume a: "X\<sharp>(domain Γ)" have "¬(∃T.(X,T)∈(set Γ))" using a proof (induct Γ) case (Cons Y Γ) thus "¬ (∃T.(X,T)∈set(Y#Γ))" by (simp only: fresh_domain_cons, auto simp add: fresh_atm) qed (simp) } ultimately show "T=S" by auto qed section {* Size and Capture-Avoiding Substitution for Types *} consts size_ty :: "ty => nat" nominal_primrec "size_ty (Tvar X) = 1" "size_ty (Top) = 1" "size_ty (T1 -> T2) = (size_ty T1) + (size_ty T2) + 1" "X\<sharp>T1 ==> size_ty (∀[X<:T1].T2) = (size_ty T1) + (size_ty T2) + 1" apply (finite_guess)+ apply (rule TrueI)+ apply (simp add: fresh_nat) apply (fresh_guess)+ done consts subst_ty :: "tyvrs => ty => ty => ty" syntax subst_ty_syn :: "ty => tyvrs => ty => ty" ("_[_:=_]ty" [100,100,100] 100) translations "T1[Y:=T2]ty" \<rightleftharpoons> "subst_ty Y T2 T1" nominal_primrec "(Tvar X)[Y:=T]ty= (if X=Y then T else (Tvar X))" "(Top)[Y:=T]ty = Top" "(T1 -> T2)[Y:=T]ty = (T1[Y:=T]ty) -> (T2[Y:=T]ty)" "[|X\<sharp>(Y,T); X\<sharp>T1|] ==> (∀[X<:T1].T2)[Y:=T]ty = (∀[X<:(T1[Y:=T]ty)].(T2[Y:=T]ty))" apply (finite_guess)+ apply (rule TrueI)+ apply (simp add: abs_fresh) apply (fresh_guess)+ done consts subst_tyc :: "ty_context => tyvrs => ty => ty_context" ("_[_:=_]tyc" [100,100,100] 100) primrec "([])[Y:=T]tyc= []" "(XT#Γ)[Y:=T]tyc = (fst XT,(snd XT)[Y:=T]ty)#(Γ[Y:=T]tyc)" section {* Subtyping-Relation *} text {* The definition for the subtyping-relation follows quite closely what is written in the POPLmark-paper, except for the premises dealing with well-formed contexts and the freshness constraint @{term "X\<sharp>Γ"} in the @{text "S_Forall"}-rule. (The freshness constraint is specific to the \emph{nominal approach}. Note, however, that the constraint does \emph{not} make the subtyping-relation ``partial"\ldots because we work over $\alpha$-equivalence classes.) *} inductive subtype_of :: "ty_context => ty => ty => bool" ("_\<turnstile>_<:_" [100,100,100] 100) where S_Top[intro]: "[|\<turnstile> Γ ok; S closed_in Γ|] ==> Γ \<turnstile> S <: Top" | S_Var[intro]: "[|(X,S) ∈ set Γ; Γ \<turnstile> S <: T|] ==> Γ \<turnstile> (Tvar X) <: T" | S_Refl[intro]: "[|\<turnstile> Γ ok; X ∈ domain Γ|]==> Γ \<turnstile> Tvar X <: Tvar X" | S_Arrow[intro]: "[|Γ \<turnstile> T1 <: S1; Γ \<turnstile> S2 <: T2|] ==> Γ \<turnstile> (S1 -> S2) <: (T1 -> T2)" | S_Forall[intro]: "[|Γ \<turnstile> T1 <: S1; X\<sharp>Γ; ((X,T1)#Γ) \<turnstile> S2 <: T2|] ==> Γ \<turnstile> ∀[X<:S1].S2 <: ∀[X<:T1].T2" lemma subtype_implies_ok: fixes X::"tyvrs" assumes a: "Γ \<turnstile> S <: T" shows "\<turnstile> Γ ok" using a by (induct) (auto) lemma subtype_implies_closed: assumes a: "Γ \<turnstile> S <: T" shows "S closed_in Γ ∧ T closed_in Γ" using a proof (induct) case (S_Top Γ S) have "Top closed_in Γ" by (simp add: closed_in_def ty.supp) moreover have "S closed_in Γ" by fact ultimately show "S closed_in Γ ∧ Top closed_in Γ" by simp next case (S_Var X S Γ T) have "(X,S)∈set Γ" by fact hence "X ∈ domain Γ" by (rule domain_inclusion) hence "(Tvar X) closed_in Γ" by (simp add: closed_in_def ty.supp supp_atm) moreover have "S closed_in Γ ∧ T closed_in Γ" by fact hence "T closed_in Γ" by force ultimately show "(Tvar X) closed_in Γ ∧ T closed_in Γ" by simp qed (auto simp add: closed_in_def ty.supp supp_atm abs_supp) lemma subtype_implies_fresh: fixes X::"tyvrs" assumes a1: "Γ \<turnstile> S <: T" and a2: "X\<sharp>Γ" shows "X\<sharp>S ∧ X\<sharp>T" proof - from a1 have "\<turnstile> Γ ok" by (rule subtype_implies_ok) with a2 have "X\<sharp>domain(Γ)" by (simp add: fresh_domain) moreover from a1 have "S closed_in Γ ∧ T closed_in Γ" by (rule subtype_implies_closed) hence "supp S ⊆ ((supp (domain Γ))::tyvrs set)" and "supp T ⊆ ((supp (domain Γ))::tyvrs set)" by (simp_all add: domain_supp closed_in_def) ultimately show "X\<sharp>S ∧ X\<sharp>T" by (force simp add: supp_prod fresh_def) qed equivariance subtype_of nominal_inductive subtype_of by (simp_all add: abs_fresh subtype_implies_fresh) thm subtype_of.strong_induct section {* Reflexivity of Subtyping *} lemma subtype_reflexivity: assumes a: "\<turnstile> Γ ok" and b: "T closed_in Γ" shows "Γ \<turnstile> T <: T" using a b proof(nominal_induct T avoiding: Γ rule: ty.induct) case (Forall X T1 T2) have ih_T1: "!!Γ. [|\<turnstile> Γ ok; T1 closed_in Γ|] ==> Γ \<turnstile> T1 <: T1" by fact have ih_T2: "!!Γ. [|\<turnstile> Γ ok; T2 closed_in Γ|] ==> Γ \<turnstile> T2 <: T2" by fact have fresh_cond: "X\<sharp>Γ" by fact hence fresh_domain: "X\<sharp>(domain Γ)" by (simp add: fresh_domain) have "(∀[X<:T2].T1) closed_in Γ" by fact hence closedT2: "T2 closed_in Γ" and closedT1: "T1 closed_in ((X,T2)#Γ)" by (auto simp add: closed_in_def ty.supp abs_supp) have ok: "\<turnstile> Γ ok" by fact hence ok': "\<turnstile> ((X,T2)#Γ) ok" using closedT2 fresh_domain by simp have "Γ \<turnstile> T2 <: T2" using ih_T2 closedT2 ok by simp moreover have "((X,T2)#Γ) \<turnstile> T1 <: T1" using ih_T1 closedT1 ok' by simp ultimately show "Γ \<turnstile> ∀[X<:T2].T1 <: ∀[X<:T2].T1" using fresh_cond by (simp add: subtype_of.S_Forall) qed (auto simp add: closed_in_def ty.supp supp_atm) lemma subtype_reflexivity_semiautomated: assumes a: "\<turnstile> Γ ok" and b: "T closed_in Γ" shows "Γ \<turnstile> T <: T" using a b apply(nominal_induct T avoiding: Γ rule: ty.induct) apply(auto simp add: ty.supp abs_supp supp_atm closed_in_def) --{* Too bad that this instantiation cannot be found automatically by \isakeyword{auto}; \isakeyword{blast} would find it if we had not used an explicit definition for @{text "closed_in_def"}. *} apply(drule_tac x="(tyvrs, ty2)#Γ" in meta_spec) apply(force dest: fresh_domain simp add: closed_in_def) done section {* Weakening *} text {* In order to prove weakening we introduce the notion of a type-context extending another. This generalization seems to make the proof for weakening to be smoother than if we had strictly adhered to the version in the POPLmark-paper. *} constdefs extends :: "ty_context => ty_context => bool" ("_ extends _" [100,100] 100) "Δ extends Γ ≡ ∀X Q. (X,Q)∈set Γ --> (X,Q)∈set Δ" lemma extends_domain: assumes a: "Δ extends Γ" shows "domain Γ ⊆ domain Δ" using a apply (auto simp add: extends_def) apply (drule domain_existence) apply (force simp add: domain_inclusion) done lemma extends_closed: assumes a1: "T closed_in Γ" and a2: "Δ extends Γ" shows "T closed_in Δ" using a1 a2 by (auto dest: extends_domain simp add: closed_in_def) lemma extends_memb: assumes a: "Δ extends Γ" and b: "(X,T) ∈ set Γ" shows "(X,T) ∈ set Δ" using a b by (simp add: extends_def) lemma weakening: assumes a: "Γ \<turnstile> S <: T" and b: "\<turnstile> Δ ok" and c: "Δ extends Γ" shows "Δ \<turnstile> S <: T" using a b c proof (nominal_induct Γ S T avoiding: Δ rule: subtype_of.strong_induct) case (S_Top Γ S) have lh_drv_prem: "S closed_in Γ" by fact have "\<turnstile> Δ ok" by fact moreover have "Δ extends Γ" by fact hence "S closed_in Δ" using lh_drv_prem by (simp only: extends_closed) ultimately show "Δ \<turnstile> S <: Top" by force next case (S_Var X S Γ T) have lh_drv_prem: "(X,S) ∈ set Γ" by fact have ih: "!!Δ. \<turnstile> Δ ok ==> Δ extends Γ ==> Δ \<turnstile> S <: T" by fact have ok: "\<turnstile> Δ ok" by fact have extends: "Δ extends Γ" by fact have "(X,S) ∈ set Δ" using lh_drv_prem extends by (simp only: extends_memb) moreover have "Δ \<turnstile> S <: T" using ok extends ih by simp ultimately show "Δ \<turnstile> Tvar X <: T" using ok by force next case (S_Refl Γ X) have lh_drv_prem: "X ∈ domain Γ" by fact have "\<turnstile> Δ ok" by fact moreover have "Δ extends Γ" by fact hence "X ∈ domain Δ" using lh_drv_prem by (force dest: extends_domain) ultimately show "Δ \<turnstile> Tvar X <: Tvar X" by force next case (S_Arrow Γ T1 S1 S2 T2) thus "Δ \<turnstile> S1 -> S2 <: T1 -> T2" by blast next case (S_Forall Γ T1 S1 X S2 T2) have fresh_cond: "X\<sharp>Δ" by fact hence fresh_domain: "X\<sharp>(domain Δ)" by (simp add: fresh_domain) have ih1: "!!Δ. \<turnstile> Δ ok ==> Δ extends Γ ==> Δ \<turnstile> T1 <: S1" by fact have ih2: "!!Δ. \<turnstile> Δ ok ==> Δ extends ((X,T1)#Γ) ==> Δ \<turnstile> S2 <: T2" by fact have lh_drv_prem: "Γ \<turnstile> T1 <: S1" by fact hence closedT1: "T1 closed_in Γ" by (simp add: subtype_implies_closed) have ok: "\<turnstile> Δ ok" by fact have ext: "Δ extends Γ" by fact have "T1 closed_in Δ" using ext closedT1 by (simp only: extends_closed) hence "\<turnstile> ((X,T1)#Δ) ok" using fresh_domain ok by force moreover have "((X,T1)#Δ) extends ((X,T1)#Γ)" using ext by (force simp add: extends_def) ultimately have "((X,T1)#Δ) \<turnstile> S2 <: T2" using ih2 by simp moreover have "Δ \<turnstile> T1 <: S1" using ok ext ih1 by simp ultimately show "Δ \<turnstile> ∀[X<:S1].S2 <: ∀[X<:T1].T2" using ok by (force intro: S_Forall) qed text {* In fact all ``non-binding" cases can be solved automatically: *} lemma weakening_more_automated: assumes a: "Γ \<turnstile> S <: T" and b: "\<turnstile> Δ ok" and c: "Δ extends Γ" shows "Δ \<turnstile> S <: T" using a b c proof (nominal_induct Γ S T avoiding: Δ rule: subtype_of.strong_induct) case (S_Forall Γ T1 S1 X S2 T2) have fresh_cond: "X\<sharp>Δ" by fact hence fresh_domain: "X\<sharp>(domain Δ)" by (simp add: fresh_domain) have ih1: "!!Δ. \<turnstile> Δ ok ==> Δ extends Γ ==> Δ \<turnstile> T1 <: S1" by fact have ih2: "!!Δ. \<turnstile> Δ ok ==> Δ extends ((X,T1)#Γ) ==> Δ \<turnstile> S2 <: T2" by fact have lh_drv_prem: "Γ \<turnstile> T1 <: S1" by fact hence closedT1: "T1 closed_in Γ" by (simp add: subtype_implies_closed) have ok: "\<turnstile> Δ ok" by fact have ext: "Δ extends Γ" by fact have "T1 closed_in Δ" using ext closedT1 by (simp only: extends_closed) hence "\<turnstile> ((X,T1)#Δ) ok" using fresh_domain ok by force moreover have "((X,T1)#Δ) extends ((X,T1)#Γ)" using ext by (force simp add: extends_def) ultimately have "((X,T1)#Δ) \<turnstile> S2 <: T2" using ih2 by simp moreover have "Δ \<turnstile> T1 <: S1" using ok ext ih1 by simp ultimately show "Δ \<turnstile> ∀[X<:S1].S2 <: ∀[X<:T1].T2" using ok by (force intro: S_Forall) qed (blast intro: extends_closed extends_memb dest: extends_domain)+ section {* Transitivity and Narrowing *} text {* Some inversion lemmas that are needed in the transitivity and narrowing proof.*} lemma S_TopE: assumes a: "Γ \<turnstile> Top <: T" shows "T = Top" using a by (cases, auto) lemma S_ArrowE_left: assumes a: "Γ \<turnstile> S1 -> S2 <: T" shows "T = Top ∨ (∃T1 T2. T = T1 -> T2 ∧ Γ \<turnstile> T1 <: S1 ∧ Γ \<turnstile> S2 <: T2)" using a by (cases, auto simp add: ty.inject) lemma S_ForallE_left: shows "[|Γ \<turnstile> ∀[X<:S1].S2 <: T; X\<sharp>Γ; X\<sharp>S1|] ==> T = Top ∨ (∃T1 T2. T = ∀[X<:T1].T2 ∧ Γ \<turnstile> T1 <: S1 ∧ ((X,T1)#Γ) \<turnstile> S2 <: T2)" apply(frule subtype_implies_ok) apply(ind_cases "Γ \<turnstile> ∀[X<:S1].S2 <: T") apply(auto simp add: ty.inject alpha) apply(rule_tac x="[(X,Xa)]•T2" in exI) apply(rule conjI) apply(rule sym) apply(rule pt_bij2[OF pt_tyvrs_inst, OF at_tyvrs_inst]) apply(rule pt_tyvrs3) apply(simp) apply(rule at_ds5[OF at_tyvrs_inst]) apply(rule conjI) apply(simp add: pt_fresh_left[OF pt_tyvrs_inst, OF at_tyvrs_inst] calc_atm) apply(drule_tac Γ="((Xa,T1)#Γ)" in subtype_implies_closed)+ apply(simp add: closed_in_def) apply(drule fresh_domain)+ apply(simp add: fresh_def) apply(subgoal_tac "X ∉ (insert Xa (domain Γ))")(*A*) apply(force) (*A*)apply(simp add: at_fin_set_supp[OF at_tyvrs_inst, OF finite_domain]) (* 2nd conjunct *)apply(frule_tac X="X" in subtype_implies_fresh) apply(assumption) apply(drule_tac X="Xa" in subtype_implies_fresh) apply(assumption) apply(simp add: fresh_prod) apply(drule_tac pi="[(X,Xa)]" in subtype_of.eqvt(2)) apply(simp add: calc_atm) apply(simp add: pt_fresh_fresh[OF pt_tyvrs_inst, OF at_tyvrs_inst]) done text {* Next we prove the transitivity and narrowing for the subtyping-relation. The POPLmark-paper says the following: \begin{quote} \begin{lemma}[Transitivity and Narrowing] \ \begin{enumerate} \item If @{term "Γ \<turnstile> S<:Q"} and @{term "Γ \<turnstile> Q<:T"}, then @{term "Γ \<turnstile> S<:T"}. \item If @{text "Γ,X<:Q,Δ \<turnstile> M<:N"} and @{term "Γ \<turnstile> P<:Q"} then @{text "Γ,X<:P,Δ \<turnstile> M<:N"}. \end{enumerate} \end{lemma} The two parts are proved simultaneously, by induction on the size of @{term "Q"}. The argument for part (2) assumes that part (1) has been established already for the @{term "Q"} in question; part (1) uses part (2) only for strictly smaller @{term "Q"}. \end{quote} For the induction on the size of @{term "Q"}, we use the induction-rule @{text "measure_induct_rule"}: \begin{center} @{thm measure_induct_rule[of "size_ty",no_vars]} \end{center} That means in order to show a property @{term "P a"} for all @{term "a"}, the induct-rule requires to prove that for all @{term x} @{term "P x"} holds using the assumption that for all @{term y} whose size is strictly smaller than that of @{term x} the property @{term "P y"} holds. *} lemma shows trans: "Γ\<turnstile>S<:Q ==> Γ\<turnstile>Q<:T ==> Γ\<turnstile>S<:T" and narrow: "(Δ@[(X,Q)]@Γ)\<turnstile>M<:N ==> Γ\<turnstile>P<:Q ==> (Δ@[(X,P)]@Γ)\<turnstile>M<:N" proof (induct Q arbitrary: Γ S T Δ X P M N taking: "size_ty" rule: measure_induct_rule) case (less Q) --{* \begin{minipage}[t]{0.9\textwidth} First we mention the induction hypotheses of the outer induction for later reference:\end{minipage}*} have IH_trans: "!!Q' Γ S T. [|size_ty Q' < size_ty Q; Γ\<turnstile>S<:Q'; Γ\<turnstile>Q'<:T|] ==> Γ\<turnstile>S<:T" by fact have IH_narrow: "!!Q' Δ Γ X M N P. [|size_ty Q' < size_ty Q; (Δ@[(X,Q')]@Γ)\<turnstile>M<:N; Γ\<turnstile>P<:Q'|] ==> (Δ@[(X,P)]@Γ)\<turnstile>M<:N" by fact --{* \begin{minipage}[t]{0.9\textwidth} We proceed with the transitivity proof as an auxiliary lemma, because it needs to be referenced in the narrowing proof.\end{minipage}*} have transitivity_aux: "!!Γ S T. [|Γ \<turnstile> S <: Q; Γ \<turnstile> Q <: T|] ==> Γ \<turnstile> S <: T" proof - fix Γ' S' T assume "Γ' \<turnstile> S' <: Q" --{* left-hand derivation *} and "Γ' \<turnstile> Q <: T" --{* right-hand derivation *} thus "Γ' \<turnstile> S' <: T" proof (nominal_induct Γ' S' Q≡Q rule: subtype_of.strong_induct) case (S_Top Γ S) --{* \begin{minipage}[t]{0.9\textwidth} In this case the left-hand derivation is @{term "Γ \<turnstile> S <: Top"}, giving us @{term "\<turnstile> Γ ok"} and @{term "S closed_in Γ"}. This case is straightforward, because the right-hand derivation must be of the form @{term "Γ \<turnstile> Top <: Top"} giving us the equation @{term "T = Top"}.\end{minipage}*} hence rh_drv: "Γ \<turnstile> Top <: T" by simp hence T_inst: "T = Top" by (simp add: S_TopE) have "\<turnstile> Γ ok" and "S closed_in Γ" by fact+ hence "Γ \<turnstile> S <: Top" by (simp add: subtype_of.S_Top) thus "Γ \<turnstile> S <: T" using T_inst by simp next case (S_Var Y U Γ) -- {* \begin{minipage}[t]{0.9\textwidth} In this case the left-hand derivation is @{term "Γ \<turnstile> Tvar Y <: Q"} with @{term "S = Tvar Y"}. We have therefore @{term "(Y,U)"} is in @{term "Γ"} and by inner induction hypothesis that @{term "Γ \<turnstile> U <: T"}. By @{text "S_Var"} follows @{term "Γ \<turnstile> Tvar Y <: T"}.\end{minipage}*} hence IH_inner: "Γ \<turnstile> U <: T" by simp have "(Y,U) ∈ set Γ" by fact with IH_inner show "Γ \<turnstile> Tvar Y <: T" by (simp add: subtype_of.S_Var) next case (S_Refl Γ X) --{* \begin{minipage}[t]{0.9\textwidth} In this case the left-hand derivation is @{term "Γ\<turnstile>(Tvar X) <: (Tvar X)"} with @{term "Q=Tvar X"}. The goal then follows immediately from the right-hand derivation.\end{minipage}*} thus "Γ \<turnstile> Tvar X <: T" by simp next case (S_Arrow Γ Q1 S1 S2 Q2) --{* \begin{minipage}[t]{0.9\textwidth} In this case the left-hand derivation is @{term "Γ \<turnstile> S1 -> S2 <: Q1 -> Q2"} with @{term "S1->S2=S"} and @{term "Q1->Q2=Q"}. We know that the @{text "size_ty"} of @{term Q1} and @{term Q2} is smaller than that of @{term Q}; so we can apply the outer induction hypotheses for @{term Q1} and @{term Q2}. We also have the sub-derivations @{term "Γ\<turnstile>Q1<:S1"} and @{term "Γ\<turnstile>S2<:Q2"}. The right-hand derivation is @{term "Γ \<turnstile> Q1 -> Q2 <: T"}. There exist types @{text "T1,T2"} such that @{term "T=Top ∨ T=T1->T2"}. The @{term "Top"}-case is straightforward once we know @{term "(S1 -> S2) closed_in Γ"} and @{term "\<turnstile> Γ ok"}. In the other case we have the sub-derivations @{term "Γ\<turnstile>T1<:Q1"} and @{term "Γ\<turnstile>Q2<:T2"}. Using the outer induction hypothesis for transitivity we can derive @{term "Γ\<turnstile>T1<:S1"} and @{term "Γ\<turnstile>S2<:T2"}. By rule @{text "S_Arrow"} follows @{term "Γ \<turnstile> S1 -> S2 <: T1 -> T2"}, which is @{term "Γ \<turnstile> S1 -> S2 <: T"}.\end{minipage}*} hence rh_drv: "Γ \<turnstile> Q1 -> Q2 <: T" by simp from `Q1 -> Q2 = Q` have Q12_less: "size_ty Q1 < size_ty Q" "size_ty Q2 < size_ty Q" by auto have lh_drv_prm1: "Γ \<turnstile> Q1 <: S1" by fact have lh_drv_prm2: "Γ \<turnstile> S2 <: Q2" by fact from rh_drv have "T=Top ∨ (∃T1 T2. T=T1->T2 ∧ Γ\<turnstile>T1<:Q1 ∧ Γ\<turnstile>Q2<:T2)" by (simp add: S_ArrowE_left) moreover have "S1 closed_in Γ" and "S2 closed_in Γ" using lh_drv_prm1 lh_drv_prm2 by (simp_all add: subtype_implies_closed) hence "(S1 -> S2) closed_in Γ" by (simp add: closed_in_def ty.supp) moreover have "\<turnstile> Γ ok" using rh_drv by (rule subtype_implies_ok) moreover { assume "∃T1 T2. T=T1->T2 ∧ Γ\<turnstile>T1<:Q1 ∧ Γ\<turnstile>Q2<:T2" then obtain T1 T2 where T_inst: "T = T1 -> T2" and rh_drv_prm1: "Γ \<turnstile> T1 <: Q1" and rh_drv_prm2: "Γ \<turnstile> Q2 <: T2" by force from IH_trans[of "Q1"] have "Γ \<turnstile> T1 <: S1" using Q12_less rh_drv_prm1 lh_drv_prm1 by simp moreover from IH_trans[of "Q2"] have "Γ \<turnstile> S2 <: T2" using Q12_less rh_drv_prm2 lh_drv_prm2 by simp ultimately have "Γ \<turnstile> S1 -> S2 <: T1 -> T2" by (simp add: subtype_of.S_Arrow) hence "Γ \<turnstile> S1 -> S2 <: T" using T_inst by simp } ultimately show "Γ \<turnstile> S1 -> S2 <: T" by blast next case (S_Forall Γ Q1 S1 X S2 Q2) --{* \begin{minipage}[t]{0.9\textwidth} In this case the left-hand derivation is @{text "Γ\<turnstile>∀[X<:S1].S2 <: ∀[X<:Q1].Q2"} with @{text "∀[X<:S1].S2=S"} and @{text "∀[X<:Q1].Q2=Q"}. We therefore have the sub-derivations @{term "Γ\<turnstile>Q1<:S1"} and @{term "((X,Q1)#Γ)\<turnstile>S2<:Q2"}. Since @{term "X"} is a binder, we assume that it is sufficiently fresh; in particular we have the freshness conditions @{term "X\<sharp>Γ"} and @{term "X\<sharp>Q1"} (these assumptions are provided by the strong induction-rule @{text "subtype_of_induct"}). We know that the @{text "size_ty"} of @{term Q1} and @{term Q2} is smaller than that of @{term Q}; so we can apply the outer induction hypotheses for @{term Q1} and @{term Q2}. The right-hand derivation is @{text "Γ \<turnstile> ∀[X<:Q1].Q2 <: T"}. Since @{term "X\<sharp>Γ"} and @{term "X\<sharp>Q1"} there exists types @{text "T1,T2"} such that @{text "T=Top ∨ T=∀[X<:T1].T2"}. The @{term "Top"}-case is straightforward once we know @{text "(∀[X<:S1].S2) closed_in Γ"} and @{term "\<turnstile> Γ ok"}. In the other case we have the sub-derivations @{term "Γ\<turnstile>T1<:Q1"} and @{term "((X,T1)#Γ)\<turnstile>Q2<:T2"}. Using the outer induction hypothesis for transitivity we can derive @{term "Γ\<turnstile>T1<:S1"}. From the outer induction for narrowing we get @{term "((X,T1)#Γ) \<turnstile> S2 <: Q2"} and then using again induction for transitivity we obtain @{term "((X,T1)#Γ) \<turnstile> S2 <: T2"}. By rule @{text "S_Forall"} and the freshness condition @{term "X\<sharp>Γ"} follows @{text "Γ \<turnstile> ∀[X<:S1].S2 <: ∀[X<:T1].T2"}, which is @{text "Γ \<turnstile> ∀[X<:S1].S2 <: T"}. \end{minipage}*} hence rh_drv: "Γ \<turnstile> ∀[X<:Q1].Q2 <: T" by simp have lh_drv_prm1: "Γ \<turnstile> Q1 <: S1" by fact have lh_drv_prm2: "((X,Q1)#Γ) \<turnstile> S2 <: Q2" by fact have "X\<sharp>Γ" by fact then have fresh_cond: "X\<sharp>Γ" "X\<sharp>Q1" using lh_drv_prm1 by (simp_all add: subtype_implies_fresh) from `∀[X<:Q1].Q2 = Q` have Q12_less: "size_ty Q1 < size_ty Q" "size_ty Q2 < size_ty Q " using fresh_cond by auto from rh_drv have "T=Top ∨ (∃T1 T2. T=∀[X<:T1].T2 ∧ Γ\<turnstile>T1<:Q1 ∧ ((X,T1)#Γ)\<turnstile>Q2<:T2)" using fresh_cond by (simp add: S_ForallE_left) moreover have "S1 closed_in Γ" and "S2 closed_in ((X,Q1)#Γ)" using lh_drv_prm1 lh_drv_prm2 by (simp_all add: subtype_implies_closed) hence "(∀[X<:S1].S2) closed_in Γ" by (force simp add: closed_in_def ty.supp abs_supp) moreover have "\<turnstile> Γ ok" using rh_drv by (rule subtype_implies_ok) moreover { assume "∃T1 T2. T=∀[X<:T1].T2 ∧ Γ\<turnstile>T1<:Q1 ∧ ((X,T1)#Γ)\<turnstile>Q2<:T2" then obtain T1 T2 where T_inst: "T = ∀[X<:T1].T2" and rh_drv_prm1: "Γ \<turnstile> T1 <: Q1" and rh_drv_prm2:"((X,T1)#Γ) \<turnstile> Q2 <: T2" by force from IH_trans[of "Q1"] have "Γ \<turnstile> T1 <: S1" using lh_drv_prm1 rh_drv_prm1 Q12_less by blast moreover from IH_narrow[of "Q1" "[]"] have "((X,T1)#Γ) \<turnstile> S2 <: Q2" using Q12_less lh_drv_prm2 rh_drv_prm1 by simp with IH_trans[of "Q2"] have "((X,T1)#Γ) \<turnstile> S2 <: T2" using Q12_less rh_drv_prm2 by simp ultimately have "Γ \<turnstile> ∀[X<:S1].S2 <: ∀[X<:T1].T2" using fresh_cond by (simp add: subtype_of.S_Forall) hence "Γ \<turnstile> ∀[X<:S1].S2 <: T" using T_inst by simp } ultimately show "Γ \<turnstile> ∀[X<:S1].S2 <: T" by blast qed qed { --{* The transitivity proof is now by the auxiliary lemma. *} case 1 have "Γ \<turnstile> S <: Q" and "Γ \<turnstile> Q <: T" by fact+ thus "Γ \<turnstile> S <: T" by (rule transitivity_aux) next --{* The narrowing proof proceeds by an induction over @{term "(Δ@[(X,Q)]@Γ) \<turnstile> M <: N"}. *} case 2 have "(Δ@[(X,Q)]@Γ) \<turnstile> M <: N" --{* left-hand derivation *} and "Γ \<turnstile> P<:Q" by fact+ --{* right-hand derivation *} thus "(Δ@[(X,P)]@Γ) \<turnstile> M <: N" proof (nominal_induct Γ≡"Δ@[(X,Q)]@Γ" M N avoiding: Δ Γ X rule: subtype_of.strong_induct) case (S_Top _ S Δ Γ X) --{* \begin{minipage}[t]{0.9\textwidth} In this case the left-hand derivation is @{term "(Δ@[(X,Q)]@Γ) \<turnstile> S <: Top"}. We show that the context @{term "Δ@[(X,P)]@Γ"} is ok and that @{term S} is closed in @{term "Δ@[(X,P)]@Γ"}. Then we can apply the @{text "S_Top"}-rule.\end{minipage}*} hence lh_drv_prm1: "\<turnstile> (Δ@[(X,Q)]@Γ) ok" and lh_drv_prm2: "S closed_in (Δ@[(X,Q)]@Γ)" by simp_all have rh_drv: "Γ \<turnstile> P <: Q" by fact hence "P closed_in Γ" by (simp add: subtype_implies_closed) with lh_drv_prm1 have "\<turnstile> (Δ@[(X,P)]@Γ) ok" by (simp add: replace_type) moreover from lh_drv_prm2 have "S closed_in (Δ@[(X,P)]@Γ)" by (simp add: closed_in_def domain_append) ultimately show "(Δ@[(X,P)]@Γ) \<turnstile> S <: Top" by (simp add: subtype_of.S_Top) next case (S_Var Y S _ N Δ Γ X) --{* \begin{minipage}[t]{0.9\textwidth} In this case the left-hand derivation is @{term "(Δ@[(X,Q)]@Γ) \<turnstile> Tvar Y <: N"} and by inner induction hypothesis we have @{term "(Δ@[(X,P)]@Γ) \<turnstile> S <: N"}. We therefore know that the contexts @{term "Δ@[(X,Q)]@Γ"} and @{term "Δ@[(X,P)]@Γ"} are ok, and that @{term "(Y,S)"} is in @{term "Δ@[(X,Q)]@Γ"}. We need to show that @{term "(Δ@[(X,P)]@Γ) \<turnstile> Tvar Y <: N"} holds. In case @{term "X≠Y"} we know that @{term "(Y,S)"} is in @{term "Δ@[(X,P)]@Γ"} and can use the inner induction hypothesis and rule @{text "S_Var"} to conclude. In case @{term "X=Y"} we can infer that @{term "S=Q"}; moreover we have that @{term "(Δ@[(X,P)]@Γ) extends Γ"} and therefore by @{text "weakening"} that @{term "(Δ@[(X,P)]@Γ) \<turnstile> P <: Q"} holds. By transitivity we obtain then @{term "(Δ@[(X,P)]@Γ) \<turnstile> P <: N"} and can conclude by applying rule @{text "S_Var"}.\end{minipage}*} hence IH_inner: "(Δ@[(X,P)]@Γ) \<turnstile> S <: N" and lh_drv_prm: "(Y,S) ∈ set (Δ@[(X,Q)]@Γ)" and rh_drv: "Γ \<turnstile> P<:Q" and okQ: "\<turnstile> (Δ@[(X,Q)]@Γ) ok" by (simp_all add: subtype_implies_ok) hence okP: "\<turnstile> (Δ@[(X,P)]@Γ) ok" by (simp add: subtype_implies_ok) show "(Δ@[(X,P)]@Γ) \<turnstile> Tvar Y <: N" proof (cases "X=Y") case False have "X≠Y" by fact hence "(Y,S)∈set (Δ@[(X,P)]@Γ)" using lh_drv_prm by simp with IH_inner show "(Δ@[(X,P)]@Γ) \<turnstile> Tvar Y <: N" by (simp add: subtype_of.S_Var) next case True have membXQ: "(X,Q)∈set (Δ@[(X,Q)]@Γ)" by simp have membXP: "(X,P)∈set (Δ@[(X,P)]@Γ)" by simp have eq: "X=Y" by fact hence "S=Q" using okQ lh_drv_prm membXQ by (simp only: uniqueness_of_ctxt) hence "(Δ@[(X,P)]@Γ) \<turnstile> Q <: N" using IH_inner by simp moreover have "(Δ@[(X,P)]@Γ) extends Γ" by (simp add: extends_def) hence "(Δ@[(X,P)]@Γ) \<turnstile> P <: Q" using rh_drv okP by (simp only: weakening) ultimately have "(Δ@[(X,P)]@Γ) \<turnstile> P <: N" by (simp add: transitivity_aux) thus "(Δ@[(X,P)]@Γ) \<turnstile> Tvar Y <: N" using membXP eq by (simp only: subtype_of.S_Var) qed next case (S_Refl _ Y Δ Γ X) --{* \begin{minipage}[t]{0.9\textwidth} In this case the left-hand derivation is @{term "(Δ@[(X,Q)]@Γ) \<turnstile> Tvar Y <: Tvar Y"} and we therefore know that @{term "Δ@[(X,Q)]@Γ"} is ok and that @{term "Y"} is in the domain of @{term "Δ@[(X,Q)]@Γ"}. We therefore know that @{term "Δ@[(X,P)]@Γ"} is ok and that @{term Y} is in the domain of @{term "Δ@[(X,P)]@Γ"}. We can conclude by applying rule @{text "S_Refl"}.\end{minipage}*} hence lh_drv_prm1: "\<turnstile> (Δ@[(X,Q)]@Γ) ok" and lh_drv_prm2: "Y ∈ domain (Δ@[(X,Q)]@Γ)" by simp_all have "Γ \<turnstile> P <: Q" by fact hence "P closed_in Γ" by (simp add: subtype_implies_closed) with lh_drv_prm1 have "\<turnstile> (Δ@[(X,P)]@Γ) ok" by (simp add: replace_type) moreover from lh_drv_prm2 have "Y ∈ domain (Δ@[(X,P)]@Γ)" by (simp add: domain_append) ultimately show "(Δ@[(X,P)]@Γ) \<turnstile> Tvar Y <: Tvar Y" by (simp add: subtype_of.S_Refl) next case (S_Arrow _ S1 Q1 Q2 S2 Δ Γ X) --{* \begin{minipage}[t]{0.9\textwidth} In this case the left-hand derivation is @{term "(Δ@[(X,Q)]@Γ) \<turnstile> Q1 -> Q2 <: S1 -> S2"} and the proof is trivial.\end{minipage}*} thus "(Δ@[(X,P)]@Γ) \<turnstile> Q1 -> Q2 <: S1 -> S2" by blast next case (S_Forall _ T1 S1 Y S2 T2 Δ Γ X) --{* \begin{minipage}[t]{0.9\textwidth} In this case the left-hand derivation is @{text "(Δ@[(X,Q)]@Γ) \<turnstile> ∀[Y<:S1].S2 <: ∀[Y<:T1].T2"} and therfore we know that the binder @{term Y} is fresh for @{term "Δ@[(X,Q)]@Γ"}. By the inner induction hypothesis we have that @{term "(Δ@[(X,P)]@Γ) \<turnstile> T1 <: S1"} and @{term "((Y,T1)#Δ@[(X,P)]@Γ) \<turnstile> S2 <: T2"}. Since @{term P} is a subtype of @{term Q} we can infer that @{term Y} is fresh for @{term P} and thus also fresh for @{term "Δ@[(X,P)]@Γ"}. We can then conclude by applying rule @{text "S_Forall"}. \end{minipage}*} hence IH_inner1: "(Δ@[(X,P)]@Γ) \<turnstile> T1 <: S1" and IH_inner2: "((Y,T1)#Δ@[(X,P)]@Γ) \<turnstile> S2 <: T2" and lh_drv_prm: "Y\<sharp>(Δ@[(X,Q)]@Γ)" by force+ have rh_drv: "Γ \<turnstile> P <: Q" by fact hence "Y\<sharp>P" using lh_drv_prm by (simp only: fresh_list_append subtype_implies_fresh) hence "Y\<sharp>(Δ@[(X,P)]@Γ)" using lh_drv_prm by (simp add: fresh_list_append fresh_list_cons fresh_prod) with IH_inner1 IH_inner2 show "(Δ@[(X,P)]@Γ) \<turnstile> ∀[Y<:S1].S2 <: ∀[Y<:T1].T2" by (simp add: subtype_of.S_Forall) qed } qed end
lemma alpha_illustration(1):
Forall X (Tvar X) T = Forall Y (Tvar Y) T
and alpha_illustration(2):
Lam x (Var x) T = Lam y (Var y) T
lemma domain_eqvt(1):
pi • domain Γ = domain (pi • Γ)
and domain_eqvt(2):
pi' • domain Γ = domain (pi' • Γ)
lemma finite_domain:
finite (domain Γ)
lemma domain_supp:
supp (domain Γ) = domain Γ
lemma domain_inclusion:
(X, T) ∈ set Γ ==> X ∈ domain Γ
lemma domain_existence:
X ∈ domain Γ ==> ∃T. (X, T) ∈ set Γ
lemma domain_append:
domain (Γ @ Δ) = domain Γ ∪ domain Δ
lemma fresh_domain_cons:
X \<sharp> domain (Y # Γ) = (X \<sharp> fst Y ∧ X \<sharp> domain Γ)
lemma fresh_domain:
X \<sharp> Γ ==> X \<sharp> domain Γ
lemma closed_in_eqvt:
S closed_in Γ ==> (pi • S) closed_in (pi • Γ)
lemma ty_vrs_prm_simp:
pi • S = S
lemma ty_context_vrs_prm_simp:
pi • Γ = Γ
lemma closed_in_eqvt':
S closed_in Γ ==> (pi • S) closed_in (pi • Γ)
lemma validE:
\<turnstile> ((X, T) # Γ) ok
==> \<turnstile> Γ ok ∧ X \<sharp> domain Γ ∧ T closed_in Γ
lemma validE_append:
\<turnstile> (Δ @ Γ) ok ==> \<turnstile> Γ ok
lemma replace_type:
[| \<turnstile> (Δ @ (X, T) # Γ) ok; S closed_in Γ |]
==> \<turnstile> (Δ @ (X, S) # Γ) ok
lemma uniqueness_of_ctxt:
[| \<turnstile> Γ ok; (X, T) ∈ set Γ; (X, S) ∈ set Γ |] ==> T = S
lemma subtype_implies_ok:
Γ\<turnstile>S<:T ==> \<turnstile> Γ ok
lemma subtype_implies_closed:
Γ\<turnstile>S<:T ==> S closed_in Γ ∧ T closed_in Γ
lemma subtype_implies_fresh:
[| Γ\<turnstile>S<:T; X \<sharp> Γ |] ==> X \<sharp> S ∧ X \<sharp> T
lemma subtype_reflexivity:
[| \<turnstile> Γ ok; T closed_in Γ |] ==> Γ\<turnstile>T<:T
lemma subtype_reflexivity_semiautomated:
[| \<turnstile> Γ ok; T closed_in Γ |] ==> Γ\<turnstile>T<:T
lemma extends_domain:
Δ extends Γ ==> domain Γ ⊆ domain Δ
lemma extends_closed:
[| T closed_in Γ; Δ extends Γ |] ==> T closed_in Δ
lemma extends_memb:
[| Δ extends Γ; (X, T) ∈ set Γ |] ==> (X, T) ∈ set Δ
lemma weakening:
[| Γ\<turnstile>S<:T; \<turnstile> Δ ok; Δ extends Γ |] ==> Δ\<turnstile>S<:T
lemma weakening_more_automated:
[| Γ\<turnstile>S<:T; \<turnstile> Δ ok; Δ extends Γ |] ==> Δ\<turnstile>S<:T
lemma S_TopE:
Γ\<turnstile>Top<:T ==> T = Top
lemma S_ArrowE_left:
Γ\<turnstile>S1 -> S2<:T
==> T = Top ∨ (∃T1 T2. T = T1 -> T2 ∧ Γ\<turnstile>T1<:S1 ∧ Γ\<turnstile>S2<:T2)
lemma S_ForallE_left:
[| Γ\<turnstile>Forall X S2 S1<:T; X \<sharp> Γ; X \<sharp> S1 |]
==> T = Top ∨
(∃T1 T2.
T = Forall X T2 T1 ∧
Γ\<turnstile>T1<:S1 ∧ ((X, T1) # Γ)\<turnstile>S2<:T2)
lemma trans:
[| Γ\<turnstile>S<:Q; Γ\<turnstile>Q<:T |] ==> Γ\<turnstile>S<:T
and narrow:
[| (Δ @ [(X, Q)] @ Γ)\<turnstile>M<:N; Γ\<turnstile>P<:Q |]
==> (Δ @ [(X, P)] @ Γ)\<turnstile>M<:N