# HG changeset patch # User Christian Urban # Date 1247689827 -7200 # Node ID 01da62fb4a209b57cc80900911c38960e8211308 # Parent a2a3685f61c3cb4690eb7b13982b680fc9304e51 simplified proofs diff -r a2a3685f61c3 -r 01da62fb4a20 src/HOL/Nominal/Examples/Fsub.thy --- a/src/HOL/Nominal/Examples/Fsub.thy Wed Jul 15 20:34:58 2009 +0200 +++ b/src/HOL/Nominal/Examples/Fsub.thy Wed Jul 15 22:30:27 2009 +0200 @@ -91,7 +91,7 @@ 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. *} + a @{text "dom"} of a typing-context is handy. *} nominal_primrec "tyvrs_of" :: "binding \ tyvrs set" @@ -108,16 +108,16 @@ by auto consts - "ty_domain" :: "env \ tyvrs set" + "ty_dom" :: "env \ tyvrs set" primrec - "ty_domain [] = {}" - "ty_domain (X#\) = (tyvrs_of X)\(ty_domain \)" + "ty_dom [] = {}" + "ty_dom (X#\) = (tyvrs_of X)\(ty_dom \)" consts - "trm_domain" :: "env \ vrs set" + "trm_dom" :: "env \ vrs set" primrec - "trm_domain [] = {}" - "trm_domain (X#\) = (vrs_of X)\(trm_domain \)" + "trm_dom [] = {}" + "trm_dom (X#\) = (vrs_of X)\(trm_dom \)" lemma vrs_of_eqvt[eqvt]: fixes pi ::"tyvrs prm" @@ -128,13 +128,13 @@ and "pi'\(vrs_of x) = vrs_of (pi'\x)" by (nominal_induct x rule: binding.strong_induct) (simp_all add: tyvrs_of.simps eqvts) -lemma domains_eqvt[eqvt]: +lemma doms_eqvt[eqvt]: fixes pi::"tyvrs prm" and pi'::"vrs prm" - shows "pi \(ty_domain \) = ty_domain (pi\\)" - and "pi'\(ty_domain \) = ty_domain (pi'\\)" - and "pi \(trm_domain \) = trm_domain (pi\\)" - and "pi'\(trm_domain \) = trm_domain (pi'\\)" + shows "pi \(ty_dom \) = ty_dom (pi\\)" + and "pi'\(ty_dom \) = ty_dom (pi'\\)" + and "pi \(trm_dom \) = trm_dom (pi\\)" + and "pi'\(trm_dom \) = trm_dom (pi'\\)" by (induct \) (simp_all add: eqvts) lemma finite_vrs: @@ -142,19 +142,19 @@ and "finite (vrs_of x)" by (nominal_induct rule:binding.strong_induct, auto) -lemma finite_domains: - shows "finite (ty_domain \)" - and "finite (trm_domain \)" +lemma finite_doms: + shows "finite (ty_dom \)" + and "finite (trm_dom \)" by (induct \, auto simp add: finite_vrs) -lemma ty_domain_supp: - shows "(supp (ty_domain \)) = (ty_domain \)" - and "(supp (trm_domain \)) = (trm_domain \)" -by (simp only: at_fin_set_supp at_tyvrs_inst at_vrs_inst finite_domains)+ +lemma ty_dom_supp: + shows "(supp (ty_dom \)) = (ty_dom \)" + and "(supp (trm_dom \)) = (trm_dom \)" +by (simp only: at_fin_set_supp at_tyvrs_inst at_vrs_inst finite_doms)+ -lemma ty_domain_inclusion: +lemma ty_dom_inclusion: assumes a: "(TVarB X T)\set \" - shows "X\(ty_domain \)" + shows "X\(ty_dom \)" using a by (induct \, auto) lemma ty_binding_existence: @@ -163,8 +163,8 @@ using assms by (nominal_induct a rule: binding.strong_induct, auto) -lemma ty_domain_existence: - assumes a: "X\(ty_domain \)" +lemma ty_dom_existence: + assumes a: "X\(ty_dom \)" shows "\T.(TVarB X T)\set \" using a apply (induct \, auto) @@ -173,9 +173,9 @@ apply (auto simp add: ty_binding_existence) done -lemma domains_append: - shows "ty_domain (\@\) = ((ty_domain \) \ (ty_domain \))" - and "trm_domain (\@\) = ((trm_domain \) \ (trm_domain \))" +lemma doms_append: + shows "ty_dom (\@\) = ((ty_dom \) \ (ty_dom \))" + and "trm_dom (\@\) = ((trm_dom \) \ (trm_dom \))" by (induct \, auto) lemma ty_vrs_prm_simp: @@ -184,15 +184,15 @@ shows "pi\S = S" by (induct S rule: ty.induct) (auto simp add: calc_atm) -lemma fresh_ty_domain_cons: +lemma fresh_ty_dom_cons: fixes X::"tyvrs" - shows "X\(ty_domain (Y#\)) = (X\(tyvrs_of Y) \ X\(ty_domain \))" + shows "X\(ty_dom (Y#\)) = (X\(tyvrs_of Y) \ X\(ty_dom \))" apply (nominal_induct rule:binding.strong_induct) apply (auto) apply (simp add: fresh_def supp_def eqvts) - apply (simp add: fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_domains) + apply (simp add: fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_doms) apply (simp add: fresh_def supp_def eqvts) - apply (simp add: fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_domains)+ + apply (simp add: fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_doms)+ done lemma tyvrs_fresh: @@ -206,26 +206,26 @@ apply (fresh_guess)+ done -lemma fresh_domain: +lemma fresh_dom: fixes X::"tyvrs" assumes a: "X\\" - shows "X\(ty_domain \)" + shows "X\(ty_dom \)" using a apply(induct \) apply(simp add: fresh_set_empty) -apply(simp only: fresh_ty_domain_cons) +apply(simp only: fresh_ty_dom_cons) apply(auto simp add: fresh_prod fresh_list_cons tyvrs_fresh) done text {* Not all lists of type @{typ "env"} are well-formed. One condition requires that in @{term "TVarB X S#\"} all free variables of @{term "S"} must be - in the @{term "ty_domain"} of @{term "\"}, that is @{term "S"} must be @{text "closed"} + in the @{term "ty_dom"} 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 \ env \ bool" ("_ closed'_in _" [100,100] 100) - "S closed_in \ \ (supp S)\(ty_domain \)" + "S closed_in \ \ (supp S)\(ty_dom \)" lemma closed_in_eqvt[eqvt]: fixes pi::"tyvrs prm" @@ -251,10 +251,10 @@ shows "x \ T" by (simp add: fresh_def supp_def ty_vrs_prm_simp) -lemma ty_domain_vrs_prm_simp: +lemma ty_dom_vrs_prm_simp: fixes pi::"vrs prm" and \::"env" - shows "(ty_domain (pi\\)) = (ty_domain \)" + shows "(ty_dom (pi\\)) = (ty_dom \)" apply(induct \) apply (simp add: eqvts) apply(simp add: tyvrs_vrs_prm_simp) @@ -265,7 +265,7 @@ assumes a: "S closed_in \" shows "(pi\S) closed_in (pi\\)" using a -by (simp add: closed_in_def ty_domain_vrs_prm_simp ty_vrs_prm_simp) +by (simp add: closed_in_def ty_dom_vrs_prm_simp ty_vrs_prm_simp) lemma fresh_vrs_of: fixes x::"vrs" @@ -273,16 +273,16 @@ by (nominal_induct b rule: binding.strong_induct) (simp_all add: fresh_singleton [OF pt_vrs_inst at_vrs_inst] fresh_set_empty ty_vrs_fresh fresh_atm) -lemma fresh_trm_domain: +lemma fresh_trm_dom: fixes x::"vrs" - shows "x\ trm_domain \ = x\\" + shows "x\ trm_dom \ = x\\" by (induct \) (simp_all add: fresh_set_empty fresh_list_cons fresh_fin_union [OF pt_vrs_inst at_vrs_inst fs_vrs_inst] - finite_domains finite_vrs fresh_vrs_of fresh_list_nil) + finite_doms finite_vrs fresh_vrs_of fresh_list_nil) -lemma closed_in_fresh: "(X::tyvrs) \ ty_domain \ \ T closed_in \ \ X \ T" - by (auto simp add: closed_in_def fresh_def ty_domain_supp) +lemma closed_in_fresh: "(X::tyvrs) \ ty_dom \ \ T closed_in \ \ X \ T" + by (auto simp add: closed_in_def fresh_def ty_dom_supp) text {* Now validity of a context is a straightforward inductive definition. *} @@ -290,15 +290,18 @@ valid_rel :: "env \ bool" ("\ _ ok" [100] 100) where valid_nil[simp]: "\ [] ok" -| valid_consT[simp]: "\\ \ ok; X\(ty_domain \); T closed_in \\ \ \ (TVarB X T#\) ok" -| valid_cons [simp]: "\\ \ ok; x\(trm_domain \); T closed_in \\ \ \ (VarB x T#\) ok" +| valid_consT[simp]: "\\ \ ok; X\(ty_dom \); T closed_in \\ \ \ (TVarB X T#\) ok" +| valid_cons [simp]: "\\ \ ok; x\(trm_dom \); T closed_in \\ \ \ (VarB x T#\) ok" equivariance valid_rel declare binding.inject [simp add] declare trm.inject [simp add] -inductive_cases validE[elim]: "\ (TVarB X T#\) ok" "\ (VarB x T#\) ok" "\ (b#\) ok" +inductive_cases validE[elim]: + "\ (TVarB X T#\) ok" + "\ (VarB x T#\) ok" + "\ (b#\) ok" declare binding.inject [simp del] declare trm.inject [simp del] @@ -321,12 +324,12 @@ using a b proof(induct \) case Nil - then show ?case by (auto elim: validE intro: valid_cons simp add: domains_append closed_in_def) + then show ?case by (auto elim: validE intro: valid_cons simp add: doms_append closed_in_def) next case (Cons a \') then show ?case by (nominal_induct a rule:binding.strong_induct) - (auto elim: validE intro!: valid_cons simp add: domains_append closed_in_def) + (auto elim: validE intro!: valid_cons simp add: doms_append closed_in_def) qed text {* Well-formed contexts have a unique type-binding for a type-variable. *} @@ -342,14 +345,14 @@ case (valid_consT \ X' T') moreover { fix \'::"env" - assume a: "X'\(ty_domain \')" + assume a: "X'\(ty_dom \')" have "\(\T.(TVarB X' T)\(set \'))" using a proof (induct \') case (Cons Y \') thus "\ (\T.(TVarB X' T)\set(Y#\'))" - by (simp add: fresh_ty_domain_cons + by (simp add: fresh_ty_dom_cons fresh_fin_union[OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] - finite_vrs finite_domains, + finite_vrs finite_doms, auto simp add: fresh_atm fresh_singleton [OF pt_tyvrs_inst at_tyvrs_inst]) qed (simp) } @@ -367,13 +370,13 @@ case (valid_cons \ x' T') moreover { fix \'::"env" - assume a: "x'\(trm_domain \')" + assume a: "x'\(trm_dom \')" have "\(\T.(VarB x' T)\(set \'))" using a proof (induct \') case (Cons y \') thus "\ (\T.(VarB x' T)\set(y#\'))" by (simp add: fresh_fin_union[OF pt_vrs_inst at_vrs_inst fs_vrs_inst] - finite_vrs finite_domains, + finite_vrs finite_doms, auto simp add: fresh_atm fresh_singleton [OF pt_vrs_inst at_vrs_inst]) qed (simp) } @@ -401,7 +404,7 @@ "(Tvar X)[Y \ T]\<^sub>\ = (if X=Y then T else Tvar X)" | "(Top)[Y \ T]\<^sub>\ = Top" | "(T\<^isub>1 \ T\<^isub>2)[Y \ T]\<^sub>\ = T\<^isub>1[Y \ T]\<^sub>\ \ T\<^isub>2[Y \ T]\<^sub>\" -| "\X\(Y,T); X\T\<^isub>1\ \ (\X<:T\<^isub>1. T\<^isub>2)[Y \ T]\<^sub>\ = (\X<:T\<^isub>1[Y \ T]\<^sub>\. T\<^isub>2[Y \ T]\<^sub>\)" +| "X\(Y,T,T\<^isub>1) \ (\X<:T\<^isub>1. T\<^isub>2)[Y \ T]\<^sub>\ = (\X<:T\<^isub>1[Y \ T]\<^sub>\. T\<^isub>2[Y \ T]\<^sub>\)" apply (finite_guess)+ apply (rule TrueI)+ apply (simp add: abs_fresh) @@ -424,8 +427,8 @@ lemma type_subst_fresh: fixes X::"tyvrs" - assumes "X \ T" and "X \ P" - shows "X \ T[Y \ P]\<^sub>\" + assumes "X\T" and "X\P" + shows "X\T[Y \ P]\<^sub>\" using assms by (nominal_induct T avoiding: X Y P rule:ty.strong_induct) (auto simp add: abs_fresh) @@ -437,17 +440,18 @@ by (nominal_induct T avoiding: X T' rule: ty.strong_induct) (auto simp add: fresh_atm abs_fresh fresh_nat) -lemma type_subst_identity: "X \ T \ T[X \ U]\<^sub>\ = T" +lemma type_subst_identity: + "X\T \ T[X \ U]\<^sub>\ = T" by (nominal_induct T avoiding: X U rule: ty.strong_induct) (simp_all add: fresh_atm abs_fresh) lemma type_substitution_lemma: - "X \ Y \ X \ L \ M[X \ N]\<^sub>\[Y \ L]\<^sub>\ = M[Y \ L]\<^sub>\[X \ N[Y \ L]\<^sub>\]\<^sub>\" + "X \ Y \ X\L \ M[X \ N]\<^sub>\[Y \ L]\<^sub>\ = M[Y \ L]\<^sub>\[X \ N[Y \ L]\<^sub>\]\<^sub>\" by (nominal_induct M avoiding: X Y N L rule: ty.strong_induct) (auto simp add: type_subst_fresh type_subst_identity) lemma type_subst_rename: - "Y \ T \ ([(Y, X)] \ T)[Y \ U]\<^sub>\ = T[X \ U]\<^sub>\" + "Y\T \ ([(Y,X)]\T)[Y \ U]\<^sub>\ = T[X \ U]\<^sub>\" by (nominal_induct T avoiding: X Y U rule: ty.strong_induct) (simp_all add: fresh_atm calc_atm abs_fresh fresh_aux) @@ -460,15 +464,15 @@ lemma binding_subst_fresh: fixes X::"tyvrs" - assumes "X \ a" - and "X \ P" - shows "X \ a[Y \ P]\<^sub>b" + assumes "X\a" + and "X\P" + shows "X\a[Y \ P]\<^sub>b" using assms by (nominal_induct a rule: binding.strong_induct) (auto simp add: type_subst_fresh) lemma binding_subst_identity: - shows "X \ B \ B[X \ U]\<^sub>b = B" + shows "X\B \ B[X \ U]\<^sub>b = B" by (induct B rule: binding.induct) (simp_all add: fresh_atm type_subst_identity) @@ -481,9 +485,9 @@ lemma ctxt_subst_fresh': fixes X::"tyvrs" - assumes "X \ \" - and "X \ P" - shows "X \ \[Y \ P]\<^sub>e" + assumes "X\\" + and "X\P" + shows "X\\[Y \ P]\<^sub>e" using assms by (induct \) (auto simp add: fresh_list_cons binding_subst_fresh) @@ -494,7 +498,7 @@ lemma ctxt_subst_mem_VarB: "VarB x T \ set \ \ VarB x (T[Y \ U]\<^sub>\) \ set (\[Y \ U]\<^sub>e)" by (induct \) auto -lemma ctxt_subst_identity: "X \ \ \ \[X \ U]\<^sub>e = \" +lemma ctxt_subst_identity: "X\\ \ \[X \ U]\<^sub>e = \" by (induct \) (simp_all add: fresh_list_cons binding_subst_identity) lemma ctxt_subst_append: "(\ @ \)[X \ T]\<^sub>e = \[X \ T]\<^sub>e @ \[X \ T]\<^sub>e" @@ -515,11 +519,13 @@ done lemma subst_trm_fresh_tyvar: - "(X::tyvrs) \ t \ X \ u \ X \ t[x \ u]" + fixes X::"tyvrs" + shows "X\t \ X\u \ X\t[x \ u]" by (nominal_induct t avoiding: x u rule: trm.strong_induct) (auto simp add: trm.fresh abs_fresh) -lemma subst_trm_fresh_var: "x \ u \ x \ t[x \ u]" +lemma subst_trm_fresh_var: + "x\u \ x\t[x \ u]" by (nominal_induct t avoiding: x u rule: trm.strong_induct) (simp_all add: abs_fresh fresh_atm ty_vrs_fresh) @@ -538,7 +544,7 @@ (perm_simp add: fresh_left)+ lemma subst_trm_rename: - "y \ t \ ([(y, x)] \ t)[y \ u] = t[x \ u]" + "y\t \ ([(y, x)] \ t)[y \ u] = t[x \ u]" by (nominal_induct t avoiding: x y u rule: trm.strong_induct) (simp_all add: fresh_atm calc_atm abs_fresh fresh_aux ty_vrs_fresh perm_fresh_fresh) @@ -558,12 +564,13 @@ done lemma subst_trm_ty_fresh: - "(X::tyvrs) \ t \ X \ T \ X \ t[Y \\<^sub>\ T]" + fixes X::"tyvrs" + shows "X\t \ X\T \ X\t[Y \\<^sub>\ T]" by (nominal_induct t avoiding: Y T rule: trm.strong_induct) (auto simp add: abs_fresh type_subst_fresh) lemma subst_trm_ty_fresh': - "X \ T \ X \ t[X \\<^sub>\ T]" + "X\T \ X\t[X \\<^sub>\ T]" by (nominal_induct t avoiding: X T rule: trm.strong_induct) (simp_all add: abs_fresh fresh_type_subst_fresh fresh_atm) @@ -582,7 +589,7 @@ (perm_simp add: fresh_left subst_eqvt')+ lemma subst_trm_ty_rename: - "Y \ t \ ([(Y, X)] \ t)[Y \\<^sub>\ U] = t[X \\<^sub>\ U]" + "Y\t \ ([(Y, X)] \ t)[Y \\<^sub>\ U] = t[X \\<^sub>\ U]" by (nominal_induct t avoiding: X Y U rule: trm.strong_induct) (simp_all add: fresh_atm calc_atm abs_fresh fresh_aux type_subst_rename) @@ -599,7 +606,7 @@ subtype_of :: "env \ ty \ ty \ bool" ("_\_<:_" [100,100,100] 100) where SA_Top[intro]: "\\ \ ok; S closed_in \\ \ \ \ S <: Top" -| SA_refl_TVar[intro]: "\\ \ ok; X \ ty_domain \\\ \ \ Tvar X <: Tvar X" +| SA_refl_TVar[intro]: "\\ \ ok; X \ ty_dom \\\ \ \ Tvar X <: Tvar X" | SA_trans_TVar[intro]: "\(TVarB X S) \ set \; \ \ S <: T\ \ \ \ (Tvar X) <: T" | SA_arrow[intro]: "\\ \ T\<^isub>1 <: S\<^isub>1; \ \ S\<^isub>2 <: T\<^isub>2\ \ \ \ (S\<^isub>1 \ S\<^isub>2) <: (T\<^isub>1 \ T\<^isub>2)" | SA_all[intro]: "\\ \ T\<^isub>1 <: S\<^isub>1; ((TVarB X T\<^isub>1)#\) \ S\<^isub>2 <: T\<^isub>2\ \ \ \ (\X<:S\<^isub>1. S\<^isub>2) <: (\X<:T\<^isub>1. T\<^isub>2)" @@ -623,7 +630,7 @@ next case (SA_trans_TVar X S \ T) have "(TVarB X S)\set \" by fact - hence "X \ ty_domain \" by (rule ty_domain_inclusion) + hence "X \ ty_dom \" by (rule ty_dom_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 @@ -638,23 +645,23 @@ shows "X\S \ X\T" proof - from a1 have "\ \ ok" by (rule subtype_implies_ok) - with a2 have "X\ty_domain(\)" by (simp add: fresh_domain) + with a2 have "X\ty_dom(\)" by (simp add: fresh_dom) moreover from a1 have "S closed_in \ \ T closed_in \" by (rule subtype_implies_closed) - hence "supp S \ ((supp (ty_domain \))::tyvrs set)" - and "supp T \ ((supp (ty_domain \))::tyvrs set)" by (simp_all add: ty_domain_supp closed_in_def) + hence "supp S \ ((supp (ty_dom \))::tyvrs set)" + and "supp T \ ((supp (ty_dom \))::tyvrs set)" by (simp_all add: ty_dom_supp closed_in_def) ultimately show "X\S \ X\T" by (force simp add: supp_prod fresh_def) qed -lemma valid_ty_domain_fresh: +lemma valid_ty_dom_fresh: fixes X::"tyvrs" assumes valid: "\ \ ok" - shows "X\(ty_domain \) = X\\" + shows "X\(ty_dom \) = X\\" using valid apply induct apply (simp add: fresh_list_nil fresh_set_empty) apply (simp_all add: binding.fresh fresh_list_cons - fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_domains fresh_atm) + fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_doms fresh_atm) apply (auto simp add: closed_in_fresh) done @@ -662,7 +669,7 @@ nominal_inductive subtype_of apply (simp_all add: abs_fresh) - apply (fastsimp simp add: valid_ty_domain_fresh dest: subtype_implies_ok) + apply (fastsimp simp add: valid_ty_dom_fresh dest: subtype_implies_ok) apply (force simp add: closed_in_fresh dest: subtype_implies_closed subtype_implies_ok)+ done @@ -678,12 +685,12 @@ have ih_T\<^isub>1: "\\. \\ \ ok; T\<^isub>1 closed_in \\ \ \ \ T\<^isub>1 <: T\<^isub>1" by fact have ih_T\<^isub>2: "\\. \\ \ ok; T\<^isub>2 closed_in \\ \ \ \ T\<^isub>2 <: T\<^isub>2" by fact have fresh_cond: "X\\" by fact - hence fresh_ty_domain: "X\(ty_domain \)" by (simp add: fresh_domain) + hence fresh_ty_dom: "X\(ty_dom \)" by (simp add: fresh_dom) have "(\X<:T\<^isub>2. T\<^isub>1) closed_in \" by fact hence closed\<^isub>T\<^isub>2: "T\<^isub>2 closed_in \" and closed\<^isub>T\<^isub>1: "T\<^isub>1 closed_in ((TVarB X T\<^isub>2)#\)" by (auto simp add: closed_in_def ty.supp abs_supp) have ok: "\ \ ok" by fact - hence ok': "\ ((TVarB X T\<^isub>2)#\) ok" using closed\<^isub>T\<^isub>2 fresh_ty_domain by simp + hence ok': "\ ((TVarB X T\<^isub>2)#\) ok" using closed\<^isub>T\<^isub>2 fresh_ty_dom by simp have "\ \ T\<^isub>2 <: T\<^isub>2" using ih_T\<^isub>2 closed\<^isub>T\<^isub>2 ok by simp moreover have "((TVarB X T\<^isub>2)#\) \ T\<^isub>1 <: T\<^isub>1" using ih_T\<^isub>1 closed\<^isub>T\<^isub>1 ok' by simp @@ -702,7 +709,7 @@ \isakeyword{auto}; \isakeyword{blast} would find it if we had not used an explicit definition for @{text "closed_in_def"}. *} apply(drule_tac x="(TVarB tyvrs ty2)#\" in meta_spec) -apply(force dest: fresh_domain simp add: closed_in_def) +apply(force dest: fresh_dom simp add: closed_in_def) done section {* Weakening *} @@ -715,13 +722,13 @@ extends :: "env \ env \ bool" ("_ extends _" [100,100] 100) "\ extends \ \ \X Q. (TVarB X Q)\set \ \ (TVarB X Q)\set \" -lemma extends_ty_domain: +lemma extends_ty_dom: assumes a: "\ extends \" - shows "ty_domain \ \ ty_domain \" + shows "ty_dom \ \ ty_dom \" using a apply (auto simp add: extends_def) - apply (drule ty_domain_existence) - apply (force simp add: ty_domain_inclusion) + apply (drule ty_dom_existence) + apply (force simp add: ty_dom_inclusion) done lemma extends_closed: @@ -729,7 +736,7 @@ and a2: "\ extends \" shows "T closed_in \" using a1 a2 - by (auto dest: extends_ty_domain simp add: closed_in_def) + by (auto dest: extends_ty_dom simp add: closed_in_def) lemma extends_memb: assumes a: "\ extends \" @@ -763,18 +770,18 @@ ultimately show "\ \ Tvar X <: T" using ok by force next case (SA_refl_TVar \ X) - have lh_drv_prem: "X \ ty_domain \" by fact + have lh_drv_prem: "X \ ty_dom \" by fact have "\ \ ok" by fact moreover have "\ extends \" by fact - hence "X \ ty_domain \" using lh_drv_prem by (force dest: extends_ty_domain) + hence "X \ ty_dom \" using lh_drv_prem by (force dest: extends_ty_dom) ultimately show "\ \ Tvar X <: Tvar X" by force next case (SA_arrow \ T\<^isub>1 S\<^isub>1 S\<^isub>2 T\<^isub>2) thus "\ \ S\<^isub>1 \ S\<^isub>2 <: T\<^isub>1 \ T\<^isub>2" by blast next case (SA_all \ T\<^isub>1 S\<^isub>1 X S\<^isub>2 T\<^isub>2) have fresh_cond: "X\\" by fact - hence fresh_domain: "X\(ty_domain \)" by (simp add: fresh_domain) + hence fresh_dom: "X\(ty_dom \)" by (simp add: fresh_dom) have ih\<^isub>1: "\\. \ \ ok \ \ extends \ \ \ \ T\<^isub>1 <: S\<^isub>1" by fact have ih\<^isub>2: "\\. \ \ ok \ \ extends ((TVarB X T\<^isub>1)#\) \ \ \ S\<^isub>2 <: T\<^isub>2" by fact have lh_drv_prem: "\ \ T\<^isub>1 <: S\<^isub>1" by fact @@ -782,7 +789,7 @@ have ok: "\ \ ok" by fact have ext: "\ extends \" by fact have "T\<^isub>1 closed_in \" using ext closed\<^isub>T\<^isub>1 by (simp only: extends_closed) - hence "\ ((TVarB X T\<^isub>1)#\) ok" using fresh_domain ok by force + hence "\ ((TVarB X T\<^isub>1)#\) ok" using fresh_dom ok by force moreover have "((TVarB X T\<^isub>1)#\) extends ((TVarB X T\<^isub>1)#\)" using ext by (force simp add: extends_def) ultimately have "((TVarB X T\<^isub>1)#\) \ S\<^isub>2 <: T\<^isub>2" using ih\<^isub>2 by simp @@ -802,7 +809,7 @@ proof (nominal_induct \ S T avoiding: \ rule: subtype_of.strong_induct) case (SA_all \ T\<^isub>1 S\<^isub>1 X S\<^isub>2 T\<^isub>2) have fresh_cond: "X\\" by fact - hence fresh_domain: "X\(ty_domain \)" by (simp add: fresh_domain) + hence fresh_dom: "X\(ty_dom \)" by (simp add: fresh_dom) have ih\<^isub>1: "\\. \ \ ok \ \ extends \ \ \ \ T\<^isub>1 <: S\<^isub>1" by fact have ih\<^isub>2: "\\. \ \ ok \ \ extends ((TVarB X T\<^isub>1)#\) \ \ \ S\<^isub>2 <: T\<^isub>2" by fact have lh_drv_prem: "\ \ T\<^isub>1 <: S\<^isub>1" by fact @@ -810,14 +817,14 @@ have ok: "\ \ ok" by fact have ext: "\ extends \" by fact have "T\<^isub>1 closed_in \" using ext closed\<^isub>T\<^isub>1 by (simp only: extends_closed) - hence "\ ((TVarB X T\<^isub>1)#\) ok" using fresh_domain ok by force + hence "\ ((TVarB X T\<^isub>1)#\) ok" using fresh_dom ok by force moreover have "((TVarB X T\<^isub>1)#\) extends ((TVarB X T\<^isub>1)#\)" using ext by (force simp add: extends_def) ultimately have "((TVarB X T\<^isub>1)#\) \ S\<^isub>2 <: T\<^isub>2" using ih\<^isub>2 by simp moreover have "\ \ T\<^isub>1 <: S\<^isub>1" using ok ext ih\<^isub>1 by simp ultimately show "\ \ (\X<:S\<^isub>1. S\<^isub>2) <: (\X<:T\<^isub>1. T\<^isub>2)" using ok by (force intro: SA_all) -qed (blast intro: extends_closed extends_memb dest: extends_ty_domain)+ +qed (blast intro: extends_closed extends_memb dest: extends_ty_dom)+ section {* Transitivity and Narrowing *} @@ -831,38 +838,11 @@ declare ty.inject [simp del] lemma S_ForallE_left: - shows "\\ \ (\X<:S\<^isub>1. S\<^isub>2) <: T; X\\; X\S\<^isub>1\ + shows "\\ \ (\X<:S\<^isub>1. S\<^isub>2) <: T; X\\; X\S\<^isub>1; X\T\ \ T = Top \ (\T\<^isub>1 T\<^isub>2. T = (\X<:T\<^isub>1. T\<^isub>2) \ \ \ T\<^isub>1 <: S\<^isub>1 \ ((TVarB X T\<^isub>1)#\) \ S\<^isub>2 <: T\<^isub>2)" - apply(frule subtype_implies_ok) - apply(ind_cases "\ \ (\X<:S\<^isub>1. S\<^isub>2) <: T") - apply(auto simp add: ty.inject alpha) - apply(rule_tac x="[(X,Xa)]\T\<^isub>2" 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 \="((TVarB Xa T\<^isub>1)#\)" 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 (ty_domain \))")(*A*) - apply(force) - (*A*)apply(simp add: at_fin_set_supp[OF at_tyvrs_inst, OF finite_domains(1)]) - (* 2nd conjunct *)apply(frule_tac X="X" in subtype_implies_fresh) - apply(assumption) - apply (frule_tac \="TVarB Xa T\<^isub>1 # \" in subtype_implies_ok) - apply (erule validE) - apply (simp add: valid_ty_domain_fresh) - apply(drule_tac X="Xa" in subtype_implies_fresh) - apply(assumption) - 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 +apply(erule subtype_of.strong_cases[where X="X"]) +apply(auto simp add: abs_fresh ty.inject alpha) +done text {* Next we prove the transitivity and narrowing for the subtyping-relation. The POPLmark-paper says the following: @@ -898,75 +878,38 @@ and subtype_narrow: "(\@[(TVarB X Q)]@\)\M<:N \ \\P<:Q \ (\@[(TVarB X P)]@\)\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; \\S<:Q'; \\Q'<:T\ \ \\S<:T" by fact have IH_narrow: "\Q' \ \ X M N P. \size_ty Q' < size_ty Q; (\@[(TVarB X Q')]@\)\M<:N; \\P<:Q'\ \ (\@[(TVarB X P)]@\)\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. \\ \ S <: Q; \ \ Q <: T\ \ \ \ S <: T" - proof - - fix \' S' T - assume "\' \ S' <: Q" --{* left-hand derivation *} - and "\' \ Q <: T" --{* right-hand derivation *} - thus "\' \ S' <: T" - proof (nominal_induct \' S' Q\Q rule: subtype_of.strong_induct) + + { fix \ S T + have "\\ \ S <: Q; \ \ Q <: T\ \ \ \ S <: T" + proof (induct \ S Q\Q rule: subtype_of.induct) case (SA_Top \ S) - --{* \begin{minipage}[t]{0.9\textwidth} - In this case the left-hand derivation is @{term "\ \ S <: Top"}, giving - us @{term "\ \ ok"} and @{term "S closed_in \"}. This case is straightforward, - because the right-hand derivation must be of the form @{term "\ \ Top <: Top"} - giving us the equation @{term "T = Top"}.\end{minipage}*} - hence rh_drv: "\ \ Top <: T" by simp - hence T_inst: "T = Top" by (auto elim: S_TopE) + then have rh_drv: "\ \ Top <: T" by simp + then have T_inst: "T = Top" by (auto elim: S_TopE) from `\ \ ok` and `S closed_in \` - have "\ \ S <: Top" by (simp add: subtype_of.SA_Top) - thus "\ \ S <: T" using T_inst by simp + have "\ \ S <: Top" by auto + then show "\ \ S <: T" using T_inst by simp next case (SA_trans_TVar Y U \) - -- {* \begin{minipage}[t]{0.9\textwidth} - In this case the left-hand derivation is @{term "\ \ Tvar Y <: Q"} - with @{term "S = Tvar Y"}. We have therefore @{term "(Y,U)"} - is in @{term "\"} and by inner induction hypothesis that @{term "\ \ U <: T"}. - By @{text "S_Var"} follows @{term "\ \ Tvar Y <: T"}.\end{minipage}*} - hence IH_inner: "\ \ U <: T" by simp + then have IH_inner: "\ \ U <: T" by simp have "(TVarB Y U) \ set \" by fact - with IH_inner show "\ \ Tvar Y <: T" by (simp add: subtype_of.SA_trans_TVar) + with IH_inner show "\ \ Tvar Y <: T" by auto next case (SA_refl_TVar \ X) - --{* \begin{minipage}[t]{0.9\textwidth} - In this case the left-hand derivation is @{term "\\(Tvar X) <: (Tvar X)"} with - @{term "Q=Tvar X"}. The goal then follows immediately from the right-hand - derivation.\end{minipage}*} - thus "\ \ Tvar X <: T" by simp + then show "\ \ Tvar X <: T" by simp next case (SA_arrow \ Q\<^isub>1 S\<^isub>1 S\<^isub>2 Q\<^isub>2) - --{* \begin{minipage}[t]{0.9\textwidth} - In this case the left-hand derivation is @{term "\ \ S\<^isub>1 \ S\<^isub>2 <: Q\<^isub>1 \ Q\<^isub>2"} with - @{term "S\<^isub>1\S\<^isub>2=S"} and @{term "Q\<^isub>1\Q\<^isub>2=Q"}. We know that the @{text "size_ty"} of - @{term Q\<^isub>1} and @{term Q\<^isub>2} is smaller than that of @{term Q}; - so we can apply the outer induction hypotheses for @{term Q\<^isub>1} and @{term Q\<^isub>2}. - We also have the sub-derivations @{term "\\Q\<^isub>1<:S\<^isub>1"} and @{term "\\S\<^isub>2<:Q\<^isub>2"}. - The right-hand derivation is @{term "\ \ Q\<^isub>1 \ Q\<^isub>2 <: T"}. There exist types - @{text "T\<^isub>1,T\<^isub>2"} such that @{term "T=Top \ T=T\<^isub>1\T\<^isub>2"}. The @{term "Top"}-case is - straightforward once we know @{term "(S\<^isub>1 \ S\<^isub>2) closed_in \"} and @{term "\ \ ok"}. - In the other case we have the sub-derivations @{term "\\T\<^isub>1<:Q\<^isub>1"} and @{term "\\Q\<^isub>2<:T\<^isub>2"}. - Using the outer induction hypothesis for transitivity we can derive @{term "\\T\<^isub>1<:S\<^isub>1"} - and @{term "\\S\<^isub>2<:T\<^isub>2"}. By rule @{text "S_Arrow"} follows @{term "\ \ S\<^isub>1 \ S\<^isub>2 <: T\<^isub>1 \ T\<^isub>2"}, - which is @{term "\ \ S\<^isub>1 \ S\<^isub>2 <: T\<^isub>"}.\end{minipage}*} - hence rh_drv: "\ \ Q\<^isub>1 \ Q\<^isub>2 <: T" by simp + then have rh_drv: "\ \ Q\<^isub>1 \ Q\<^isub>2 <: T" by simp from `Q\<^isub>1 \ Q\<^isub>2 = Q` have Q\<^isub>1\<^isub>2_less: "size_ty Q\<^isub>1 < size_ty Q" "size_ty Q\<^isub>2 < size_ty Q" by auto have lh_drv_prm\<^isub>1: "\ \ Q\<^isub>1 <: S\<^isub>1" by fact have lh_drv_prm\<^isub>2: "\ \ S\<^isub>2 <: Q\<^isub>2" by fact from rh_drv have "T=Top \ (\T\<^isub>1 T\<^isub>2. T=T\<^isub>1\T\<^isub>2 \ \\T\<^isub>1<:Q\<^isub>1 \ \\Q\<^isub>2<:T\<^isub>2)" - by (auto elim: S_ArrowE_left) + by (auto elim: S_ArrowE_left) moreover have "S\<^isub>1 closed_in \" and "S\<^isub>2 closed_in \" using lh_drv_prm\<^isub>1 lh_drv_prm\<^isub>2 by (simp_all add: subtype_implies_closed) @@ -974,7 +917,7 @@ moreover have "\ \ ok" using rh_drv by (rule subtype_implies_ok) moreover - { assume "\T\<^isub>1 T\<^isub>2. T=T\<^isub>1\T\<^isub>2 \ \\T\<^isub>1<:Q\<^isub>1 \ \\Q\<^isub>2<:T\<^isub>2" + { assume "\T\<^isub>1 T\<^isub>2. T = T\<^isub>1\T\<^isub>2 \ \ \ T\<^isub>1 <: Q\<^isub>1 \ \ \ Q\<^isub>2 <: T\<^isub>2" then obtain T\<^isub>1 T\<^isub>2 where T_inst: "T = T\<^isub>1 \ T\<^isub>2" and rh_drv_prm\<^isub>1: "\ \ T\<^isub>1 <: Q\<^isub>1" @@ -984,46 +927,26 @@ moreover from IH_trans[of "Q\<^isub>2"] have "\ \ S\<^isub>2 <: T\<^isub>2" using Q\<^isub>1\<^isub>2_less rh_drv_prm\<^isub>2 lh_drv_prm\<^isub>2 by simp - ultimately have "\ \ S\<^isub>1 \ S\<^isub>2 <: T\<^isub>1 \ T\<^isub>2" by (simp add: subtype_of.SA_arrow) - hence "\ \ S\<^isub>1 \ S\<^isub>2 <: T" using T_inst by simp + ultimately have "\ \ S\<^isub>1 \ S\<^isub>2 <: T\<^isub>1 \ T\<^isub>2" by auto + then have "\ \ S\<^isub>1 \ S\<^isub>2 <: T" using T_inst by simp } ultimately show "\ \ S\<^isub>1 \ S\<^isub>2 <: T" by blast next case (SA_all \ Q\<^isub>1 S\<^isub>1 X S\<^isub>2 Q\<^isub>2) - --{* \begin{minipage}[t]{0.9\textwidth} - In this case the left-hand derivation is @{term "\\(\X<:S\<^isub>1. S\<^isub>2) <: (\X<:Q\<^isub>1. Q\<^isub>2)"} with - @{term "(\X<:S\<^isub>1. S\<^isub>2)=S"} and @{term "(\X<:Q\<^isub>1. Q\<^isub>2)=Q"}. We therefore have the sub-derivations - @{term "\\Q\<^isub>1<:S\<^isub>1"} and @{term "((TVarB X Q\<^isub>1)#\)\S\<^isub>2<:Q\<^isub>2"}. Since @{term "X"} is a binder, we - assume that it is sufficiently fresh; in particular we have the freshness conditions - @{term "X\\"} and @{term "X\Q\<^isub>1"} (these assumptions are provided by the strong - induction-rule @{text "subtype_of_induct"}). We know that the @{text "size_ty"} of - @{term Q\<^isub>1} and @{term Q\<^isub>2} is smaller than that of @{term Q}; - so we can apply the outer induction hypotheses for @{term Q\<^isub>1} and @{term Q\<^isub>2}. - The right-hand derivation is @{term "\ \ (\X<:Q\<^isub>1. Q\<^isub>2) <: T"}. Since @{term "X\\"} - and @{term "X\Q\<^isub>1"} there exists types @{text "T\<^isub>1,T\<^isub>2"} such that - @{term "T=Top \ T=(\X<:T\<^isub>1. T\<^isub>2)"}. The @{term "Top"}-case is straightforward once we know - @{term "(\X<:S\<^isub>1. S\<^isub>2) closed_in \"} and @{term "\ \ ok"}. In the other case we have - the sub-derivations @{term "\\T\<^isub>1<:Q\<^isub>1"} and @{term "((TVarB X T\<^isub>1)#\)\Q\<^isub>2<:T\<^isub>2"}. Using the outer - induction hypothesis for transitivity we can derive @{term "\\T\<^isub>1<:S\<^isub>1"}. From the outer - induction for narrowing we get @{term "((TVarB X T\<^isub>1)#\) \ S\<^isub>2 <: Q\<^isub>2"} and then using again - induction for transitivity we obtain @{term "((TVarB X T\<^isub>1)#\) \ S\<^isub>2 <: T\<^isub>2"}. By rule - @{text "S_Forall"} and the freshness condition @{term "X\\"} follows - @{term "\ \ (\X<:S\<^isub>1. S\<^isub>2) <: (\X<:T\<^isub>1. T\<^isub>2)"}, which is @{term "\ \ (\X<:S\<^isub>1. S\<^isub>2) <: T\<^isub>"}. - \end{minipage}*} - hence rh_drv: "\ \ (\X<:Q\<^isub>1. Q\<^isub>2) <: T" by simp + then have rh_drv: "\ \ (\X<:Q\<^isub>1. Q\<^isub>2) <: T" by simp have lh_drv_prm\<^isub>1: "\ \ Q\<^isub>1 <: S\<^isub>1" by fact have lh_drv_prm\<^isub>2: "((TVarB X Q\<^isub>1)#\) \ S\<^isub>2 <: Q\<^isub>2" by fact - then have "X\\" by (force dest: subtype_implies_ok simp add: valid_ty_domain_fresh) - then have fresh_cond: "X\\" "X\Q\<^isub>1" using lh_drv_prm\<^isub>1 by (simp_all add: subtype_implies_fresh) - from `(\X<:Q\<^isub>1. Q\<^isub>2) = Q` - have Q\<^isub>1\<^isub>2_less: "size_ty Q\<^isub>1 < size_ty Q" "size_ty Q\<^isub>2 < size_ty Q " using fresh_cond by auto + then have "X\\" by (force dest: subtype_implies_ok simp add: valid_ty_dom_fresh) + then have fresh_cond: "X\\" "X\Q\<^isub>1" "X\T" using rh_drv lh_drv_prm\<^isub>1 + by (simp_all add: subtype_implies_fresh) from rh_drv - have "T=Top \ (\T\<^isub>1 T\<^isub>2. T=(\X<:T\<^isub>1. T\<^isub>2) \ \\T\<^isub>1<:Q\<^isub>1 \ ((TVarB X T\<^isub>1)#\)\Q\<^isub>2<:T\<^isub>2)" + have "T = Top \ + (\T\<^isub>1 T\<^isub>2. T = (\X<:T\<^isub>1. T\<^isub>2) \ \ \ T\<^isub>1 <: Q\<^isub>1 \ ((TVarB X T\<^isub>1)#\) \ Q\<^isub>2 <: T\<^isub>2)" using fresh_cond by (simp add: S_ForallE_left) moreover have "S\<^isub>1 closed_in \" and "S\<^isub>2 closed_in ((TVarB X Q\<^isub>1)#\)" using lh_drv_prm\<^isub>1 lh_drv_prm\<^isub>2 by (simp_all add: subtype_implies_closed) - hence "(\X<:S\<^isub>1. S\<^isub>2) closed_in \" by (force simp add: closed_in_def ty.supp abs_supp) + then have "(\X<:S\<^isub>1. S\<^isub>2) closed_in \" by (force simp add: closed_in_def ty.supp abs_supp) moreover have "\ \ ok" using rh_drv by (rule subtype_implies_ok) moreover @@ -1032,6 +955,9 @@ where T_inst: "T = (\X<:T\<^isub>1. T\<^isub>2)" and rh_drv_prm\<^isub>1: "\ \ T\<^isub>1 <: Q\<^isub>1" and rh_drv_prm\<^isub>2:"((TVarB X T\<^isub>1)#\) \ Q\<^isub>2 <: T\<^isub>2" by force + have "(\X<:Q\<^isub>1. Q\<^isub>2) = Q" by fact + then have Q\<^isub>1\<^isub>2_less: "size_ty Q\<^isub>1 < size_ty Q" "size_ty Q\<^isub>2 < size_ty Q" + using fresh_cond by auto from IH_trans[of "Q\<^isub>1"] have "\ \ T\<^isub>1 <: S\<^isub>1" using lh_drv_prm\<^isub>1 rh_drv_prm\<^isub>1 Q\<^isub>1\<^isub>2_less by blast moreover @@ -1045,52 +971,35 @@ } ultimately show "\ \ (\X<:S\<^isub>1. S\<^isub>2) <: T" by blast qed - qed + } note transitivity_lemma = this { --{* The transitivity proof is now by the auxiliary lemma. *} case 1 from `\ \ S <: Q` and `\ \ Q <: T` - show "\ \ S <: T" by (rule transitivity_aux) + show "\ \ S <: T" by (rule transitivity_lemma) next - --{* The narrowing proof proceeds by an induction over @{term "(\@[(TVarB X Q)]@\) \ M <: N"}. *} case 2 - from `(\@[(TVarB X Q)]@\) \ M <: N` --{* left-hand derivation *} - and `\ \ P<:Q` --{* right-hand derivation *} + from `(\@[(TVarB X Q)]@\) \ M <: N` + and `\ \ P<:Q` show "(\@[(TVarB X P)]@\) \ M <: N" - proof (nominal_induct \\"\@[(TVarB X Q)]@\" M N avoiding: \ \ X rule: subtype_of.strong_induct) - case (SA_Top _ S \ \ X) - --{* \begin{minipage}[t]{0.9\textwidth} - In this case the left-hand derivation is @{term "(\@[(TVarB X Q)]@\) \ S <: Top"}. We show - that the context @{term "\@[(TVarB X P)]@\"} is ok and that @{term S} is closed in - @{term "\@[(TVarB X P)]@\"}. Then we can apply the @{text "S_Top"}-rule.\end{minipage}*} - hence lh_drv_prm\<^isub>1: "\ (\@[(TVarB X Q)]@\) ok" + proof (induct \\"\@[(TVarB X Q)]@\" M N arbitrary: \ X \ rule: subtype_of.induct) + case (SA_Top _ S \ X \) + then have lh_drv_prm\<^isub>1: "\ (\@[(TVarB X Q)]@\) ok" and lh_drv_prm\<^isub>2: "S closed_in (\@[(TVarB X Q)]@\)" by simp_all have rh_drv: "\ \ P <: Q" by fact hence "P closed_in \" by (simp add: subtype_implies_closed) with lh_drv_prm\<^isub>1 have "\ (\@[(TVarB X P)]@\) ok" by (simp add: replace_type) moreover from lh_drv_prm\<^isub>2 have "S closed_in (\@[(TVarB X P)]@\)" - by (simp add: closed_in_def domains_append) + by (simp add: closed_in_def doms_append) ultimately show "(\@[(TVarB X P)]@\) \ S <: Top" by (simp add: subtype_of.SA_Top) next - case (SA_trans_TVar Y S _ N \ \ X) - --{* \begin{minipage}[t]{0.9\textwidth} - In this case the left-hand derivation is @{term "(\@[(TVarB X Q)]@\) \ Tvar Y <: N"} and - by inner induction hypothesis we have @{term "(\@[(TVarB X P)]@\) \ S <: N"}. We therefore - know that the contexts @{term "\@[(TVarB X Q)]@\"} and @{term "\@[(TVarB X P)]@\"} are ok, and that - @{term "(Y,S)"} is in @{term "\@[(TVarB X Q)]@\"}. We need to show that - @{term "(\@[(TVarB X P)]@\) \ Tvar Y <: N"} holds. In case @{term "X\Y"} we know that - @{term "(Y,S)"} is in @{term "\@[(TVarB 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 "(\@[(TVarB X P)]@\) extends \"} and therefore - by @{text "weakening"} that @{term "(\@[(TVarB X P)]@\) \ P <: Q"} holds. By transitivity we - obtain then @{term "(\@[(TVarB X P)]@\) \ P <: N"} and can conclude by applying rule - @{text "S_Var"}.\end{minipage}*} - hence IH_inner: "(\@[(TVarB X P)]@\) \ S <: N" + case (SA_trans_TVar Y S _ N \ X \) + then have IH_inner: "(\@[(TVarB X P)]@\) \ S <: N" and lh_drv_prm: "(TVarB Y S) \ set (\@[(TVarB X Q)]@\)" and rh_drv: "\ \ P<:Q" and ok\<^isub>Q: "\ (\@[(TVarB X Q)]@\) ok" by (simp_all add: subtype_implies_ok) - hence ok\<^isub>P: "\ (\@[(TVarB X P)]@\) ok" by (simp add: subtype_implies_ok) + then have ok\<^isub>P: "\ (\@[(TVarB X P)]@\) ok" by (simp add: subtype_implies_ok) show "(\@[(TVarB X P)]@\) \ Tvar Y <: N" proof (cases "X=Y") case False @@ -1107,48 +1016,28 @@ moreover have "(\@[(TVarB X P)]@\) extends \" by (simp add: extends_def) hence "(\@[(TVarB X P)]@\) \ P <: Q" using rh_drv ok\<^isub>P by (simp only: weakening) - ultimately have "(\@[(TVarB X P)]@\) \ P <: N" by (simp add: transitivity_aux) - thus "(\@[(TVarB X P)]@\) \ Tvar Y <: N" using memb\<^isub>X\<^isub>P eq by (simp only: subtype_of.SA_trans_TVar) + ultimately have "(\@[(TVarB X P)]@\) \ P <: N" by (simp add: transitivity_lemma) + then show "(\@[(TVarB X P)]@\) \ Tvar Y <: N" using memb\<^isub>X\<^isub>P eq by auto qed next - case (SA_refl_TVar _ Y \ \ X) - --{* \begin{minipage}[t]{0.9\textwidth} - In this case the left-hand derivation is @{term "(\@[(TVarB X Q)]@\) \ Tvar Y <: Tvar Y"} and we - therefore know that @{term "\@[(TVarB X Q)]@\"} is ok and that @{term "Y"} is in - the domain of @{term "\@[(TVarB X Q)]@\"}. We therefore know that @{term "\@[(TVarB X P)]@\"} is ok - and that @{term Y} is in the domain of @{term "\@[(TVarB X P)]@\"}. We can conclude by applying - rule @{text "S_Refl"}.\end{minipage}*} - hence lh_drv_prm\<^isub>1: "\ (\@[(TVarB X Q)]@\) ok" - and lh_drv_prm\<^isub>2: "Y \ ty_domain (\@[(TVarB X Q)]@\)" by simp_all + case (SA_refl_TVar _ Y \ X \) + then have lh_drv_prm\<^isub>1: "\ (\@[(TVarB X Q)]@\) ok" + and lh_drv_prm\<^isub>2: "Y \ ty_dom (\@[(TVarB X Q)]@\)" by simp_all have "\ \ P <: Q" by fact hence "P closed_in \" by (simp add: subtype_implies_closed) with lh_drv_prm\<^isub>1 have "\ (\@[(TVarB X P)]@\) ok" by (simp add: replace_type) moreover - from lh_drv_prm\<^isub>2 have "Y \ ty_domain (\@[(TVarB X P)]@\)" by (simp add: domains_append) + from lh_drv_prm\<^isub>2 have "Y \ ty_dom (\@[(TVarB X P)]@\)" by (simp add: doms_append) ultimately show "(\@[(TVarB X P)]@\) \ Tvar Y <: Tvar Y" by (simp add: subtype_of.SA_refl_TVar) next - case (SA_arrow _ S\<^isub>1 Q\<^isub>1 Q\<^isub>2 S\<^isub>2 \ \ X) - --{* \begin{minipage}[t]{0.9\textwidth} - In this case the left-hand derivation is @{term "(\@[(TVarB X Q)]@\) \ Q\<^isub>1 \ Q\<^isub>2 <: S\<^isub>1 \ S\<^isub>2"} - and the proof is trivial.\end{minipage}*} - thus "(\@[(TVarB X P)]@\) \ Q\<^isub>1 \ Q\<^isub>2 <: S\<^isub>1 \ S\<^isub>2" by blast + case (SA_arrow _ S\<^isub>1 Q\<^isub>1 Q\<^isub>2 S\<^isub>2 \ X \) + then show "(\@[(TVarB X P)]@\) \ Q\<^isub>1 \ Q\<^isub>2 <: S\<^isub>1 \ S\<^isub>2" by blast next - case (SA_all \' T\<^isub>1 S\<^isub>1 Y S\<^isub>2 T\<^isub>2 \ \ X) - --{* \begin{minipage}[t]{0.9\textwidth} - In this case the left-hand derivation is @{term "(\@[(TVarB X Q)]@\) \ (\Y<:S\<^isub>1. S\<^isub>2) <: (\Y<:T\<^isub>1. T\<^isub>2)"} - and therfore we know that the binder @{term Y} is fresh for @{term "\@[(TVarB X Q)]@\"}. By - the inner induction hypothesis we have that @{term "(\@[(TVarB X P)]@\) \ T\<^isub>1 <: S\<^isub>1"} and - @{term "((TVarB Y T\<^isub>1)#\@[(TVarB X P)]@\) \ S\<^isub>2 <: T\<^isub>2"}. 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 "\@[(TVarB X P)]@\"}. We can then conclude by applying rule @{text "S_Forall"}. - \end{minipage}*} - hence rh_drv: "\ \ P <: Q" - and IH_inner\<^isub>1: "(\@[(TVarB X P)]@\) \ T\<^isub>1 <: S\<^isub>1" - and "TVarB Y T\<^isub>1 # \' = ((TVarB Y T\<^isub>1)#\) @ [TVarB X Q] @ \" by auto - moreover have " \\\P<:Q; TVarB Y T\<^isub>1 # \' = ((TVarB Y T\<^isub>1)#\) @ [TVarB X Q] @ \\ \ (((TVarB Y T\<^isub>1)#\) @ [TVarB X P] @ \)\S\<^isub>2<:T\<^isub>2" by fact - ultimately have IH_inner\<^isub>2: "(((TVarB Y T\<^isub>1)#\) @ [TVarB X P] @ \)\S\<^isub>2<:T\<^isub>2" by auto - with IH_inner\<^isub>1 IH_inner\<^isub>2 - show "(\@[(TVarB X P)]@\) \ (\Y<:S\<^isub>1. S\<^isub>2) <: (\Y<:T\<^isub>1. T\<^isub>2)" by (simp add: subtype_of.SA_all) + case (SA_all _ T\<^isub>1 S\<^isub>1 Y S\<^isub>2 T\<^isub>2 \ X \) + from SA_all(2,4,5,6) + have IH_inner\<^isub>1: "(\@[(TVarB X P)]@\) \ T\<^isub>1 <: S\<^isub>1" + and IH_inner\<^isub>2: "(((TVarB Y T\<^isub>1)#\)@[(TVarB X P)]@\) \ S\<^isub>2 <: T\<^isub>2" by force+ + then show "(\@[(TVarB X P)]@\) \ (\Y<:S\<^isub>1. S\<^isub>2) <: (\Y<:T\<^isub>1. T\<^isub>2)" by auto qed } qed @@ -1163,7 +1052,7 @@ | T_Abs[intro]: "\ VarB x T\<^isub>1 # \ \ t\<^isub>2 : T\<^isub>2 \ \ \ \ (\x:T\<^isub>1. t\<^isub>2) : T\<^isub>1 \ T\<^isub>2" | T_Sub[intro]: "\ \ \ t : S; \ \ S <: T \ \ \ \ t : T" | T_TAbs[intro]:"\ TVarB X T\<^isub>1 # \ \ t\<^isub>2 : T\<^isub>2 \ \ \ \ (\X<:T\<^isub>1. t\<^isub>2) : (\X<:T\<^isub>1. T\<^isub>2)" -| T_TApp[intro]:"\ X \ (\, t\<^isub>1, T\<^isub>2); \ \ t\<^isub>1 : (\X<:T\<^isub>1\<^isub>1. T\<^isub>1\<^isub>2); \ \ T\<^isub>2 <: T\<^isub>1\<^isub>1 \ \ \ \ t\<^isub>1 \\<^sub>\ T\<^isub>2 : (T\<^isub>1\<^isub>2[X \ T\<^isub>2]\<^sub>\)" +| T_TApp[intro]:"\X\(\,t\<^isub>1,T\<^isub>2); \ \ t\<^isub>1 : (\X<:T\<^isub>1\<^isub>1. T\<^isub>1\<^isub>2); \ \ T\<^isub>2 <: T\<^isub>1\<^isub>1\ \ \ \ t\<^isub>1 \\<^sub>\ T\<^isub>2 : (T\<^isub>1\<^isub>2[X \ T\<^isub>2]\<^sub>\)" equivariance typing @@ -1189,8 +1078,8 @@ using assms by (induct, auto) nominal_inductive typing -by (auto dest!: typing_ok intro: closed_in_fresh fresh_domain type_subst_fresh - simp: abs_fresh fresh_type_subst_fresh ty_vrs_fresh valid_ty_domain_fresh fresh_trm_domain) +by (auto dest!: typing_ok intro: closed_in_fresh fresh_dom type_subst_fresh + simp: abs_fresh fresh_type_subst_fresh ty_vrs_fresh valid_ty_dom_fresh fresh_trm_dom) lemma ok_imp_VarB_closed_in: assumes ok: "\ \ ok" @@ -1200,19 +1089,19 @@ lemma tyvrs_of_subst: "tyvrs_of (B[X \ T]\<^sub>b) = tyvrs_of B" by (nominal_induct B rule: binding.strong_induct) simp_all -lemma ty_domain_subst: "ty_domain (\[X \ T]\<^sub>e) = ty_domain \" +lemma ty_dom_subst: "ty_dom (\[X \ T]\<^sub>e) = ty_dom \" by (induct \) (simp_all add: tyvrs_of_subst) lemma vrs_of_subst: "vrs_of (B[X \ T]\<^sub>b) = vrs_of B" by (nominal_induct B rule: binding.strong_induct) simp_all -lemma trm_domain_subst: "trm_domain (\[X \ T]\<^sub>e) = trm_domain \" +lemma trm_dom_subst: "trm_dom (\[X \ T]\<^sub>e) = trm_dom \" by (induct \) (simp_all add: vrs_of_subst) lemma subst_closed_in: "T closed_in (\ @ TVarB X S # \) \ U closed_in \ \ T[X \ U]\<^sub>\ closed_in (\[X \ U]\<^sub>e @ \)" apply (nominal_induct T avoiding: X U \ rule: ty.strong_induct) - apply (simp add: closed_in_def ty.supp supp_atm domains_append ty_domain_subst) + apply (simp add: closed_in_def ty.supp supp_atm doms_append ty_dom_subst) apply blast apply (simp add: closed_in_def ty.supp) apply (simp add: closed_in_def ty.supp) @@ -1220,7 +1109,7 @@ apply (drule_tac x = X in meta_spec) apply (drule_tac x = U in meta_spec) apply (drule_tac x = "(TVarB tyvrs ty2) # \" in meta_spec) - apply (simp add: domains_append ty_domain_subst) + apply (simp add: doms_append ty_dom_subst) apply blast done @@ -1323,20 +1212,20 @@ "(\x:T. t) \ t'" "(\X<:T. t) \ t'" -lemma ty_domain_cons: - shows "ty_domain (\@[VarB X Q]@\) = ty_domain (\@\)" +lemma ty_dom_cons: + shows "ty_dom (\@[VarB X Q]@\) = ty_dom (\@\)" by (induct \, auto) lemma closed_in_cons: assumes "S closed_in (\ @ VarB X Q # \)" shows "S closed_in (\@\)" -using assms ty_domain_cons closed_in_def by auto +using assms ty_dom_cons closed_in_def by auto lemma closed_in_weaken: "T closed_in (\ @ \) \ T closed_in (\ @ B # \)" - by (auto simp add: closed_in_def domains_append) + by (auto simp add: closed_in_def doms_append) lemma closed_in_weaken': "T closed_in \ \ T closed_in (\ @ \)" - by (auto simp add: closed_in_def domains_append) + by (auto simp add: closed_in_def doms_append) lemma valid_subst: assumes ok: "\ (\ @ TVarB X Q # \) ok" @@ -1350,24 +1239,24 @@ apply simp apply (rule valid_consT) apply assumption - apply (simp add: domains_append ty_domain_subst) - apply (simp add: fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_domains) + apply (simp add: doms_append ty_dom_subst) + apply (simp add: fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_doms) apply (rule_tac S=Q in subst_closed_in') - apply (simp add: closed_in_def domains_append ty_domain_subst) - apply (simp add: closed_in_def domains_append) + apply (simp add: closed_in_def doms_append ty_dom_subst) + apply (simp add: closed_in_def doms_append) apply blast apply simp apply (rule valid_cons) apply assumption - apply (simp add: domains_append trm_domain_subst) + apply (simp add: doms_append trm_dom_subst) apply (rule_tac S=Q in subst_closed_in') - apply (simp add: closed_in_def domains_append ty_domain_subst) - apply (simp add: closed_in_def domains_append) + apply (simp add: closed_in_def doms_append ty_dom_subst) + apply (simp add: closed_in_def doms_append) apply blast done -lemma ty_domain_vrs: - shows "ty_domain (G @ [VarB x Q] @ D) = ty_domain (G @ D)" +lemma ty_dom_vrs: + shows "ty_dom (G @ [VarB x Q] @ D) = ty_dom (G @ D)" by (induct G, auto) lemma valid_cons': @@ -1389,10 +1278,10 @@ case (Cons b bs) with valid_consT have "\ (bs @ \) ok" by simp - moreover from Cons and valid_consT have "X \ ty_domain (bs @ \)" - by (simp add: domains_append) + moreover from Cons and valid_consT have "X \ ty_dom (bs @ \)" + by (simp add: doms_append) moreover from Cons and valid_consT have "T closed_in (bs @ \)" - by (simp add: closed_in_def domains_append) + by (simp add: closed_in_def doms_append) ultimately have "\ (TVarB X T # bs @ \) ok" by (rule valid_rel.valid_consT) with Cons and valid_consT show ?thesis by simp @@ -1407,11 +1296,11 @@ case (Cons b bs) with valid_cons have "\ (bs @ \) ok" by simp - moreover from Cons and valid_cons have "x \ trm_domain (bs @ \)" - by (simp add: domains_append finite_domains + moreover from Cons and valid_cons have "x \ trm_dom (bs @ \)" + by (simp add: doms_append finite_doms fresh_fin_insert [OF pt_vrs_inst at_vrs_inst fs_vrs_inst]) moreover from Cons and valid_cons have "T closed_in (bs @ \)" - by (simp add: closed_in_def domains_append) + by (simp add: closed_in_def doms_append) ultimately have "\ (VarB x T # bs @ \) ok" by (rule valid_rel.valid_cons) with Cons and valid_cons show ?thesis by simp @@ -1439,10 +1328,10 @@ have "\ (VarB y T\<^isub>1 # \ @ B # \) ok" apply (rule valid_cons) apply (rule T_Abs) - apply (simp add: domains_append + apply (simp add: doms_append fresh_fin_insert [OF pt_vrs_inst at_vrs_inst fs_vrs_inst] fresh_fin_union [OF pt_vrs_inst at_vrs_inst fs_vrs_inst] - finite_domains finite_vrs fresh_vrs_of T_Abs fresh_trm_domain) + finite_doms finite_vrs fresh_vrs_of T_Abs fresh_trm_dom) apply (rule closed_in_weaken) apply (rule closed) done @@ -1467,10 +1356,10 @@ have "\ (TVarB X T\<^isub>1 # \ @ B # \) ok" apply (rule valid_consT) apply (rule T_TAbs) - apply (simp add: domains_append + apply (simp add: doms_append fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] fresh_fin_union [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] - finite_domains finite_vrs tyvrs_fresh T_TAbs fresh_domain) + finite_doms finite_vrs tyvrs_fresh T_TAbs fresh_dom) apply (rule closed_in_weaken) apply (rule closed) done @@ -1513,8 +1402,8 @@ case (SA_refl_TVar G X' G') then have "\ (G' @ VarB x Q # \) ok" by simp then have h1:"\ (G' @ \) ok" by (auto dest: valid_cons') - have "X' \ ty_domain (G' @ VarB x Q # \)" using SA_refl_TVar by auto - then have h2:"X' \ ty_domain (G' @ \)" using ty_domain_vrs by auto + have "X' \ ty_dom (G' @ VarB x Q # \)" using SA_refl_TVar by auto + then have h2:"X' \ ty_dom (G' @ \)" using ty_dom_vrs by auto show ?case using h1 h2 by auto next case (SA_all G T1 S1 X S2 T2 G') @@ -1592,7 +1481,7 @@ next case (T_TAbs X T1 G t2 T2 x u D) from `TVarB X T1 # G \ t2 : T2` have "X \ T1" - by (auto simp add: valid_ty_domain_fresh dest: typing_ok intro!: closed_in_fresh) + by (auto simp add: valid_ty_dom_fresh dest: typing_ok intro!: closed_in_fresh) with `X \ u` and T_TAbs show ?case by fastsimp next case (T_TApp X G t1 T2 T11 T12 x u D) @@ -1627,7 +1516,7 @@ assume eq: "X = Y" from eq and SA_trans_TVar have "TVarB Y Q \ set G" by simp with G_ok have QS: "Q = S" using `TVarB Y S \ set G` by (rule uniqueness_of_ctxt) - from X\_ok have "X \ ty_domain \" and "Q closed_in \" by auto + from X\_ok have "X \ ty_dom \" and "Q closed_in \" by auto then have XQ: "X \ Q" by (rule closed_in_fresh) note `\\P<:Q` moreover from ST have "\ (D[X \ P]\<^sub>e @ \) ok" by (rule subtype_implies_ok) @@ -1652,8 +1541,8 @@ with neq and ST show ?thesis by auto next assume Y: "TVarB Y S \ set \" - from X\_ok have "X \ ty_domain \" and "\ \ ok" by auto - then have "X \ \" by (simp add: valid_ty_domain_fresh) + from X\_ok have "X \ ty_dom \" and "\ \ ok" by auto + then have "X \ \" by (simp add: valid_ty_dom_fresh) with Y have "X \ S" by (induct \) (auto simp add: fresh_list_nil fresh_list_cons) with ST have "(D[X \ P]\<^sub>e @ \)\S<:T[X \ P]\<^sub>\" @@ -1677,8 +1566,8 @@ by (auto intro: subtype_reflexivity) next assume neq: "X \ Y" - with SA_refl_TVar have "Y \ ty_domain (D[X \ P]\<^sub>e @ \)" - by (simp add: ty_domain_subst domains_append) + with SA_refl_TVar have "Y \ ty_dom (D[X \ P]\<^sub>e @ \)" + by (simp add: ty_dom_subst doms_append) with neq and ok show ?thesis by auto qed next @@ -1688,8 +1577,8 @@ show ?case using subtype_of.SA_arrow h1 h2 by auto next case (SA_all G T1 S1 Y S2 T2 X P D) - then have Y: "Y \ ty_domain (D @ TVarB X Q # \)" - by (auto dest: subtype_implies_ok intro: fresh_domain) + then have Y: "Y \ ty_dom (D @ TVarB X Q # \)" + by (auto dest: subtype_implies_ok intro: fresh_dom) moreover from SA_all have "S1 closed_in (D @ TVarB X Q # \)" by (auto dest: subtype_implies_closed) ultimately have S1: "Y \ S1" by (rule closed_in_fresh) @@ -1722,8 +1611,8 @@ next assume x: "VarB x T \ set G" from T_Var have ok: "\ G ok" by (auto dest: subtype_implies_ok) - then have "X \ ty_domain G" using T_Var by (auto dest: validE_append) - with ok have "X \ G" by (simp add: valid_ty_domain_fresh) + then have "X \ ty_dom G" using T_Var by (auto dest: validE_append) + with ok have "X \ G" by (simp add: valid_ty_dom_fresh) moreover from x have "VarB x T \ set (D' @ G)" by simp then have "VarB x (T[X \ P]\<^sub>\) \ set ((D' @ G)[X \ P]\<^sub>e)" by (rule ctxt_subst_mem_VarB) @@ -1744,15 +1633,15 @@ then show ?case using substT_subtype by force next case (T_TAbs X' G' T1 t2 T2 X P D') - then have "X' \ ty_domain (D' @ TVarB X Q # G)" + then have "X' \ ty_dom (D' @ TVarB X Q # G)" and "G' closed_in (D' @ TVarB X Q # G)" by (auto dest: typing_ok) then have "X' \ G'" by (rule closed_in_fresh) with T_TAbs show ?case by force next case (T_TApp X' G' t1 T2 T11 T12 X P D') - then have "X' \ ty_domain (D' @ TVarB X Q # G)" - by (simp add: fresh_domain) + then have "X' \ ty_dom (D' @ TVarB X Q # G)" + by (simp add: fresh_dom) moreover from T_TApp have "T11 closed_in (D' @ TVarB X Q # G)" by (auto dest: subtype_implies_closed) ultimately have X': "X' \ T11" by (rule closed_in_fresh) @@ -1783,7 +1672,7 @@ then have "[(y, x)] \ (VarB y S # \) \ [(y, x)] \ [(y, x)] \ s : [(y, x)] \ T\<^isub>2" by (rule typing.eqvt) moreover from T_Abs have "y \ \" - by (auto dest!: typing_ok simp add: fresh_trm_domain) + by (auto dest!: typing_ok simp add: fresh_trm_dom) ultimately have "VarB x S # \ \ s : T\<^isub>2" using T_Abs by (perm_simp add: ty_vrs_prm_simp) with ty1 show ?case using ty2 by (rule T_Abs) @@ -1819,7 +1708,7 @@ proof (nominal_induct \ t \ "\X<:S. s" T avoiding: X U U' S arbitrary: s rule: typing.strong_induct) case (T_TAbs Y T\<^isub>1 \ t\<^isub>2 T\<^isub>2) from `TVarB Y T\<^isub>1 # \ \ t\<^isub>2 : T\<^isub>2` have Y: "Y \ \" - by (auto dest!: typing_ok simp add: valid_ty_domain_fresh) + by (auto dest!: typing_ok simp add: valid_ty_dom_fresh) from `Y \ U'` and `Y \ X` have "(\X<:U. U') = (\Y<:U. [(Y, X)] \ U')" by (simp add: ty.inject alpha' fresh_atm) @@ -1907,7 +1796,7 @@ with T_TApp have "\ \ (\X<:T\<^isub>1\<^isub>1'. t\<^isub>1\<^isub>2) : (\X<:T\<^isub>1\<^isub>1. T\<^isub>1\<^isub>2)" and "X \ \" and "X \ T\<^isub>1\<^isub>1'" by (simp_all add: trm.inject) moreover from `\\T\<^isub>2<:T\<^isub>1\<^isub>1` and `X \ \` have "X \ T\<^isub>1\<^isub>1" - by (blast intro: closed_in_fresh fresh_domain dest: subtype_implies_closed) + by (blast intro: closed_in_fresh fresh_dom dest: subtype_implies_closed) ultimately obtain S' where "TVarB X T\<^isub>1\<^isub>1 # \ \ t\<^isub>1\<^isub>2 : S'" and "(TVarB X T\<^isub>1\<^isub>1 # \) \ S' <: T\<^isub>1\<^isub>2"