| author | urbanc | 
| Thu, 12 Jun 2008 09:56:28 +0200 | |
| changeset 27162 | 8d747de5c73e | 
| parent 26966 | 071f40487734 | 
| child 29097 | 68245155eb58 | 
| permissions | -rw-r--r-- | 
| 18266 
55c201fe4c95
added an authors section (please let me know if somebody is left out or unhappy)
 urbanc parents: 
18263diff
changeset | 1 | (* $Id$ *) | 
| 
55c201fe4c95
added an authors section (please let me know if somebody is left out or unhappy)
 urbanc parents: 
18263diff
changeset | 2 | |
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 3 | (*<*) | 
| 19501 | 4 | theory Fsub | 
| 5 | imports "../Nominal" | |
| 18269 | 6 | begin | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 7 | (*>*) | 
| 18269 | 8 | |
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 9 | text{* Authors: Christian Urban,
 | 
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 10 | Benjamin Pierce, | 
| 18650 | 11 | Dimitrios Vytiniotis | 
| 12 | Stephanie Weirich and | |
| 13 | Steve Zdancewic | |
| 18266 
55c201fe4c95
added an authors section (please let me know if somebody is left out or unhappy)
 urbanc parents: 
18263diff
changeset | 14 | |
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 15 | with great help from Stefan Berghofer and Markus Wenzel. *} | 
| 18246 | 16 | |
| 18621 | 17 | section {* Types for Names, Nominal Datatype Declaration for Types and Terms *}
 | 
| 18424 | 18 | |
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 19 | text {* The main point of this solution is to use names everywhere (be they bound, 
 | 
| 18621 | 20 |   binding or free). In System \FSUB{} there are two kinds of names corresponding to 
 | 
| 21 | type-variables and to term-variables. These two kinds of names are represented in | |
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 22 |   the nominal datatype package as atom-types @{text "tyvrs"} and @{text "vrs"}: *}
 | 
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 23 | |
| 18246 | 24 | atom_decl tyvrs vrs | 
| 25 | ||
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 26 | text{* There are numerous facts that come with this declaration: for example that 
 | 
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 27 |   there are infinitely many elements in @{text "tyvrs"} and @{text "vrs"}. *}
 | 
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 28 | |
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 29 | text{* The constructors for types and terms in System \FSUB{} contain abstractions 
 | 
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 30 | over type-variables and term-variables. The nominal datatype-package uses | 
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 31 |   @{text "\<guillemotleft>\<dots>\<guillemotright>\<dots>"} to indicate where abstractions occur. *}
 | 
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 32 | |
| 18424 | 33 | nominal_datatype ty = | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 34 | Tvar "tyvrs" | 
| 18424 | 35 | | Top | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 36 |   | Arrow  "ty" "ty"          ("_ \<rightarrow> _" [100,100] 100)
 | 
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 37 | | Forall "\<guillemotleft>tyvrs\<guillemotright>ty" "ty" | 
| 18246 | 38 | |
| 18424 | 39 | nominal_datatype trm = | 
| 40 | Var "vrs" | |
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 41 | | Lam "\<guillemotleft>vrs\<guillemotright>trm" "ty" | 
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 42 | | Tabs "\<guillemotleft>tyvrs\<guillemotright>trm" "ty" | 
| 18424 | 43 | | App "trm" "trm" | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 44 | | Tapp "trm" "ty" | 
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 45 | |
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 46 | text {* To be polite to the eye, some more familiar notation is introduced. 
 | 
| 18621 | 47 | Because of the change in the order of arguments, one needs to use | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 48 | translation rules, instead of syntax annotations at the term-constructors | 
| 18650 | 49 |   as given above for @{term "Arrow"}. *}
 | 
| 18246 | 50 | |
| 51 | syntax | |
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 52 |   Forall_syn :: "tyvrs \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> ty" ("\<forall>[_<:_]._" [100,100,100] 100)
 | 
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 53 |   Lam_syn    :: "vrs \<Rightarrow> ty \<Rightarrow> trm \<Rightarrow> trm"   ("Lam [_:_]._" [100,100,100] 100)
 | 
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 54 |   Tabs_syn   :: "tyvrs \<Rightarrow> ty \<Rightarrow> trm \<Rightarrow> trm" ("Tabs [_<:_]._" [100,100,100] 100)
 | 
| 18424 | 55 | |
| 18246 | 56 | translations | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 57 | "\<forall>[X<:T\<^isub>1].T\<^isub>2" \<rightleftharpoons> "ty.Forall X T\<^isub>2 T\<^isub>1" | 
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 58 | "Lam [x:T].t" \<rightleftharpoons> "trm.Lam x t T" | 
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 59 | "Tabs [X<:T].t" \<rightleftharpoons> "trm.Tabs X t T" | 
| 18246 | 60 | |
| 18621 | 61 | text {* Again there are numerous facts that are proved automatically for @{typ "ty"} 
 | 
| 18650 | 62 |   and @{typ "trm"}: for example that the set of free variables, i.e.~the @{text "support"}, 
 | 
| 63 |   is finite. However note that nominal-datatype declarations do \emph{not} define 
 | |
| 64 | ``classical" constructor-based datatypes, but rather define $\alpha$-equivalence | |
| 18621 | 65 |   classes---we can for example show that $\alpha$-equivalent @{typ "ty"}s 
 | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 66 |   and @{typ "trm"}s are equal: *}
 | 
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 67 | |
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 68 | lemma alpha_illustration: | 
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 69 | shows "\<forall>[X<:T].(Tvar X) = \<forall>[Y<:T].(Tvar Y)" | 
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 70 | and "Lam [x:T].(Var x) = Lam [y:T].(Var y)" | 
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 71 | by (simp_all add: ty.inject trm.inject alpha calc_atm fresh_atm) | 
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 72 | |
| 18621 | 73 | section {* SubTyping Contexts *}
 | 
| 18246 | 74 | |
| 75 | types ty_context = "(tyvrs\<times>ty) list" | |
| 76 | ||
| 18650 | 77 | text {* Typing contexts are represented as lists that ``grow" on the left; we
 | 
| 18621 | 78 | thereby deviating from the convention in the POPLmark-paper. The lists contain | 
| 18650 | 79 | pairs of type-variables and types (this is sufficient for Part 1A). *} | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 80 | |
| 18628 | 81 | text {* In order to state validity-conditions for typing-contexts, the notion of
 | 
| 18621 | 82 |   a @{text "domain"} of a typing-context is handy. *}
 | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 83 | |
| 18246 | 84 | consts | 
| 85 | "domain" :: "ty_context \<Rightarrow> tyvrs set" | |
| 86 | primrec | |
| 87 |   "domain [] = {}"
 | |
| 88 |   "domain (X#\<Gamma>) = {fst X}\<union>(domain \<Gamma>)" 
 | |
| 89 | ||
| 22537 | 90 | lemma domain_eqvt[eqvt]: | 
| 18246 | 91 | fixes pi::"tyvrs prm" | 
| 22537 | 92 | and pi'::"vrs prm" | 
| 93 | shows "pi\<bullet>(domain \<Gamma>) = domain (pi\<bullet>\<Gamma>)" | |
| 94 | and "pi'\<bullet>(domain \<Gamma>) = domain (pi'\<bullet>\<Gamma>)" | |
| 22541 
c33b542394f3
the name for the collection of equivariance lemmas is now eqvts (changed from eqvt) in order to avoid clashes with eqvt-lemmas generated in nominal_inductive
 urbanc parents: 
22537diff
changeset | 95 | by (induct \<Gamma>) (simp_all add: eqvts) | 
| 18246 | 96 | |
| 97 | lemma finite_domain: | |
| 98 | shows "finite (domain \<Gamma>)" | |
| 99 | by (induct \<Gamma>, auto) | |
| 100 | ||
| 18621 | 101 | lemma domain_supp: | 
| 102 | shows "(supp (domain \<Gamma>)) = (domain \<Gamma>)" | |
| 103 | by (simp only: at_fin_set_supp at_tyvrs_inst finite_domain) | |
| 104 | ||
| 18424 | 105 | lemma domain_inclusion: | 
| 106 | assumes a: "(X,T)\<in>set \<Gamma>" | |
| 107 | shows "X\<in>(domain \<Gamma>)" | |
| 108 | using a by (induct \<Gamma>, auto) | |
| 18246 | 109 | |
| 18424 | 110 | lemma domain_existence: | 
| 111 | assumes a: "X\<in>(domain \<Gamma>)" | |
| 112 | shows "\<exists>T.(X,T)\<in>set \<Gamma>" | |
| 113 | using a by (induct \<Gamma>, auto) | |
| 18246 | 114 | |
| 115 | lemma domain_append: | |
| 116 | shows "domain (\<Gamma>@\<Delta>) = ((domain \<Gamma>) \<union> (domain \<Delta>))" | |
| 117 | by (induct \<Gamma>, auto) | |
| 118 | ||
| 18621 | 119 | lemma fresh_domain_cons: | 
| 120 | fixes X::"tyvrs" | |
| 121 | shows "X\<sharp>(domain (Y#\<Gamma>)) = (X\<sharp>(fst Y) \<and> X\<sharp>(domain \<Gamma>))" | |
| 122 | by (simp add: fresh_fin_insert pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst finite_domain) | |
| 123 | ||
| 124 | lemma fresh_domain: | |
| 125 | fixes X::"tyvrs" | |
| 126 | assumes a: "X\<sharp>\<Gamma>" | |
| 127 | shows "X\<sharp>(domain \<Gamma>)" | |
| 128 | using a | |
| 129 | apply(induct \<Gamma>) | |
| 130 | apply(simp add: fresh_set_empty) | |
| 131 | apply(simp only: fresh_domain_cons) | |
| 132 | apply(auto simp add: fresh_prod fresh_list_cons) | |
| 133 | done | |
| 134 | ||
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 135 | text {* Not all lists of type @{typ "ty_context"} are well-formed. One condition
 | 
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 136 |   requires that in @{term "(X,S)#\<Gamma>"} all free variables of @{term "S"} must be 
 | 
| 18621 | 137 |   in the @{term "domain"} of @{term "\<Gamma>"}, that is @{term "S"} must be @{text "closed"} 
 | 
| 18650 | 138 |   in @{term "\<Gamma>"}. The set of free variables of @{term "S"} is the 
 | 
| 18621 | 139 |   @{text "support"} of @{term "S"}. *}
 | 
| 18246 | 140 | |
| 141 | constdefs | |
| 142 |   "closed_in" :: "ty \<Rightarrow> ty_context \<Rightarrow> bool" ("_ closed'_in _" [100,100] 100)
 | |
| 143 | "S closed_in \<Gamma> \<equiv> (supp S)\<subseteq>(domain \<Gamma>)" | |
| 144 | ||
| 22537 | 145 | lemma closed_in_eqvt[eqvt]: | 
| 18246 | 146 | fixes pi::"tyvrs prm" | 
| 147 | assumes a: "S closed_in \<Gamma>" | |
| 148 | shows "(pi\<bullet>S) closed_in (pi\<bullet>\<Gamma>)" | |
| 149 | using a | |
| 26091 | 150 | proof - | 
| 151 | from a have "pi\<bullet>(S closed_in \<Gamma>)" by (simp add: perm_bool) | |
| 152 | then show "(pi\<bullet>S) closed_in (pi\<bullet>\<Gamma>)" by (simp add: closed_in_def eqvts) | |
| 18246 | 153 | qed | 
| 154 | ||
| 22537 | 155 | lemma ty_vrs_prm_simp: | 
| 156 | fixes pi::"vrs prm" | |
| 157 | and S::"ty" | |
| 158 | shows "pi\<bullet>S = S" | |
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26091diff
changeset | 159 | by (induct S rule: ty.induct) (auto simp add: calc_atm) | 
| 22537 | 160 | |
| 161 | lemma ty_context_vrs_prm_simp: | |
| 162 | fixes pi::"vrs prm" | |
| 163 | and \<Gamma>::"ty_context" | |
| 164 | shows "pi\<bullet>\<Gamma> = \<Gamma>" | |
| 165 | by (induct \<Gamma>) | |
| 166 | (auto simp add: calc_atm ty_vrs_prm_simp) | |
| 167 | ||
| 168 | lemma closed_in_eqvt'[eqvt]: | |
| 169 | fixes pi::"vrs prm" | |
| 170 | assumes a: "S closed_in \<Gamma>" | |
| 171 | shows "(pi\<bullet>S) closed_in (pi\<bullet>\<Gamma>)" | |
| 172 | using a | |
| 173 | by (simp add: ty_vrs_prm_simp ty_context_vrs_prm_simp) | |
| 174 | ||
| 18621 | 175 | text {* Now validity of a context is a straightforward inductive definition. *}
 | 
| 176 | ||
| 23760 | 177 | inductive | 
| 22537 | 178 |   valid_rel :: "ty_context \<Rightarrow> bool" ("\<turnstile> _ ok" [100] 100)
 | 
| 22436 | 179 | where | 
| 180 | valid_nil[simp]: "\<turnstile> [] ok" | |
| 181 | | valid_cons[simp]: "\<lbrakk>\<turnstile> \<Gamma> ok; X\<sharp>(domain \<Gamma>); T closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<turnstile> ((X,T)#\<Gamma>) ok" | |
| 18246 | 182 | |
| 22537 | 183 | equivariance valid_rel | 
| 18246 | 184 | |
| 185 | lemma validE: | |
| 186 | assumes a: "\<turnstile> ((X,T)#\<Gamma>) ok" | |
| 18621 | 187 | shows "\<turnstile> \<Gamma> ok \<and> X\<sharp>(domain \<Gamma>) \<and> T closed_in \<Gamma>" | 
| 18246 | 188 | using a by (cases, auto) | 
| 189 | ||
| 18424 | 190 | lemma validE_append: | 
| 191 | assumes a: "\<turnstile> (\<Delta>@\<Gamma>) ok" | |
| 192 | shows "\<turnstile> \<Gamma> ok" | |
| 193 | using a by (induct \<Delta>, auto dest: validE) | |
| 18246 | 194 | |
| 18424 | 195 | lemma replace_type: | 
| 196 | assumes a: "\<turnstile> (\<Delta>@(X,T)#\<Gamma>) ok" | |
| 197 | and b: "S closed_in \<Gamma>" | |
| 198 | shows "\<turnstile> (\<Delta>@(X,S)#\<Gamma>) ok" | |
| 18621 | 199 | using a b | 
| 18246 | 200 | apply(induct \<Delta>) | 
| 18621 | 201 | apply(auto dest!: validE intro!: valid_cons simp add: domain_append closed_in_def) | 
| 18246 | 202 | done | 
| 203 | ||
| 18650 | 204 | text {* Well-formed contexts have a unique type-binding for a type-variable. *} 
 | 
| 205 | ||
| 18246 | 206 | lemma uniqueness_of_ctxt: | 
| 207 | fixes \<Gamma>::"ty_context" | |
| 18412 | 208 | assumes a: "\<turnstile> \<Gamma> ok" | 
| 209 | and b: "(X,T)\<in>set \<Gamma>" | |
| 210 | and c: "(X,S)\<in>set \<Gamma>" | |
| 211 | shows "T=S" | |
| 18621 | 212 | using a b c | 
| 213 | proof (induct) | |
| 214 | case valid_nil thus "T=S" by simp | |
| 215 | next | |
| 22436 | 216 | case valid_cons | 
| 18621 | 217 | moreover | 
| 218 |   { fix \<Gamma>::"ty_context"
 | |
| 219 | assume a: "X\<sharp>(domain \<Gamma>)" | |
| 220 | have "\<not>(\<exists>T.(X,T)\<in>(set \<Gamma>))" using a | |
| 221 | proof (induct \<Gamma>) | |
| 222 | case (Cons Y \<Gamma>) | |
| 223 | thus "\<not> (\<exists>T.(X,T)\<in>set(Y#\<Gamma>))" | |
| 224 | by (simp only: fresh_domain_cons, auto simp add: fresh_atm) | |
| 225 | qed (simp) | |
| 226 | } | |
| 227 | ultimately show "T=S" by auto | |
| 228 | qed | |
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 229 | |
| 18628 | 230 | section {* Size and Capture-Avoiding Substitution for Types *}
 | 
| 18621 | 231 | |
| 21554 | 232 | consts size_ty :: "ty \<Rightarrow> nat" | 
| 20395 
9a60e3151244
added definition for size and substitution using the recursion
 urbanc parents: 
19972diff
changeset | 233 | |
| 21554 | 234 | nominal_primrec | 
| 235 | "size_ty (Tvar X) = 1" | |
| 236 | "size_ty (Top) = 1" | |
| 237 | "size_ty (T1 \<rightarrow> T2) = (size_ty T1) + (size_ty T2) + 1" | |
| 238 | "X\<sharp>T1 \<Longrightarrow> size_ty (\<forall>[X<:T1].T2) = (size_ty T1) + (size_ty T2) + 1" | |
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
21554diff
changeset | 239 | apply (finite_guess)+ | 
| 21554 | 240 | apply (rule TrueI)+ | 
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
21554diff
changeset | 241 | apply (simp add: fresh_nat) | 
| 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
21554diff
changeset | 242 | apply (fresh_guess)+ | 
| 21554 | 243 | done | 
| 20395 
9a60e3151244
added definition for size and substitution using the recursion
 urbanc parents: 
19972diff
changeset | 244 | |
| 21554 | 245 | consts subst_ty :: "tyvrs \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> ty" | 
| 20395 
9a60e3151244
added definition for size and substitution using the recursion
 urbanc parents: 
19972diff
changeset | 246 | |
| 
9a60e3151244
added definition for size and substitution using the recursion
 urbanc parents: 
19972diff
changeset | 247 | syntax | 
| 
9a60e3151244
added definition for size and substitution using the recursion
 urbanc parents: 
19972diff
changeset | 248 |  subst_ty_syn :: "ty \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> ty" ("_[_:=_]\<^isub>t\<^isub>y" [100,100,100] 100)
 | 
| 
9a60e3151244
added definition for size and substitution using the recursion
 urbanc parents: 
19972diff
changeset | 249 | |
| 
9a60e3151244
added definition for size and substitution using the recursion
 urbanc parents: 
19972diff
changeset | 250 | translations | 
| 
9a60e3151244
added definition for size and substitution using the recursion
 urbanc parents: 
19972diff
changeset | 251 | "T1[Y:=T2]\<^isub>t\<^isub>y" \<rightleftharpoons> "subst_ty Y T2 T1" | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 252 | |
| 21554 | 253 | nominal_primrec | 
| 254 | "(Tvar X)[Y:=T]\<^isub>t\<^isub>y= (if X=Y then T else (Tvar X))" | |
| 255 | "(Top)[Y:=T]\<^isub>t\<^isub>y = Top" | |
| 256 | "(T\<^isub>1 \<rightarrow> T\<^isub>2)[Y:=T]\<^isub>t\<^isub>y = (T\<^isub>1[Y:=T]\<^isub>t\<^isub>y) \<rightarrow> (T\<^isub>2[Y:=T]\<^isub>t\<^isub>y)" | |
| 257 | "\<lbrakk>X\<sharp>(Y,T); X\<sharp>T\<^isub>1\<rbrakk> \<Longrightarrow> (\<forall>[X<:T\<^isub>1].T\<^isub>2)[Y:=T]\<^isub>t\<^isub>y = (\<forall>[X<:(T\<^isub>1[Y:=T]\<^isub>t\<^isub>y)].(T\<^isub>2[Y:=T]\<^isub>t\<^isub>y))" | |
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
21554diff
changeset | 258 | apply (finite_guess)+ | 
| 21554 | 259 | apply (rule TrueI)+ | 
| 260 | apply (simp add: abs_fresh) | |
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
21554diff
changeset | 261 | apply (fresh_guess)+ | 
| 21554 | 262 | done | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 263 | |
| 22537 | 264 | consts | 
| 265 |   subst_tyc :: "ty_context \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> ty_context" ("_[_:=_]\<^isub>t\<^isub>y\<^isub>c" [100,100,100] 100)
 | |
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 266 | primrec | 
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 267 | "([])[Y:=T]\<^isub>t\<^isub>y\<^isub>c= []" | 
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 268 | "(XT#\<Gamma>)[Y:=T]\<^isub>t\<^isub>y\<^isub>c = (fst XT,(snd XT)[Y:=T]\<^isub>t\<^isub>y)#(\<Gamma>[Y:=T]\<^isub>t\<^isub>y\<^isub>c)" | 
| 18246 | 269 | |
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 270 | section {* Subtyping-Relation *}
 | 
| 18246 | 271 | |
| 18650 | 272 | text {* The definition for the subtyping-relation follows quite closely what is written 
 | 
| 273 | in the POPLmark-paper, except for the premises dealing with well-formed contexts and | |
| 274 |   the freshness constraint @{term "X\<sharp>\<Gamma>"} in the @{text "S_Forall"}-rule. (The freshness
 | |
| 275 |   constraint is specific to the \emph{nominal approach}. Note, however, that the constraint
 | |
| 276 |   does \emph{not} make the subtyping-relation ``partial"\ldots because we work over
 | |
| 277 | $\alpha$-equivalence classes.) *} | |
| 18628 | 278 | |
| 23760 | 279 | inductive | 
| 22537 | 280 |   subtype_of :: "ty_context \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool"   ("_\<turnstile>_<:_" [100,100,100] 100)
 | 
| 22436 | 281 | where | 
| 282 | S_Top[intro]: "\<lbrakk>\<turnstile> \<Gamma> ok; S closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> S <: Top" | |
| 283 | | S_Var[intro]: "\<lbrakk>(X,S) \<in> set \<Gamma>; \<Gamma> \<turnstile> S <: T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (Tvar X) <: T" | |
| 284 | | S_Refl[intro]: "\<lbrakk>\<turnstile> \<Gamma> ok; X \<in> domain \<Gamma>\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> Tvar X <: Tvar X" | |
| 285 | | S_Arrow[intro]: "\<lbrakk>\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1; \<Gamma> \<turnstile> S\<^isub>2 <: T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (S\<^isub>1 \<rightarrow> S\<^isub>2) <: (T\<^isub>1 \<rightarrow> T\<^isub>2)" | |
| 22537 | 286 | | S_Forall[intro]: "\<lbrakk>\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1; X\<sharp>\<Gamma>; ((X,T\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> \<forall>[X<:S\<^isub>1].S\<^isub>2 <: \<forall>[X<:T\<^isub>1].T\<^isub>2" | 
| 287 | ||
| 288 | lemma subtype_implies_ok: | |
| 289 | fixes X::"tyvrs" | |
| 290 | assumes a: "\<Gamma> \<turnstile> S <: T" | |
| 291 | shows "\<turnstile> \<Gamma> ok" | |
| 292 | using a by (induct) (auto) | |
| 18246 | 293 | |
| 294 | lemma subtype_implies_closed: | |
| 295 | assumes a: "\<Gamma> \<turnstile> S <: T" | |
| 296 | shows "S closed_in \<Gamma> \<and> T closed_in \<Gamma>" | |
| 297 | using a | |
| 298 | proof (induct) | |
| 22436 | 299 | case (S_Top \<Gamma> S) | 
| 18424 | 300 | have "Top closed_in \<Gamma>" by (simp add: closed_in_def ty.supp) | 
| 18246 | 301 | moreover | 
| 302 | have "S closed_in \<Gamma>" by fact | |
| 303 | ultimately show "S closed_in \<Gamma> \<and> Top closed_in \<Gamma>" by simp | |
| 304 | next | |
| 22436 | 305 | case (S_Var X S \<Gamma> T) | 
| 18246 | 306 | have "(X,S)\<in>set \<Gamma>" by fact | 
| 18424 | 307 | hence "X \<in> domain \<Gamma>" by (rule domain_inclusion) | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 308 | hence "(Tvar X) closed_in \<Gamma>" by (simp add: closed_in_def ty.supp supp_atm) | 
| 18246 | 309 | moreover | 
| 310 | have "S closed_in \<Gamma> \<and> T closed_in \<Gamma>" by fact | |
| 311 | hence "T closed_in \<Gamma>" by force | |
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 312 | ultimately show "(Tvar X) closed_in \<Gamma> \<and> T closed_in \<Gamma>" by simp | 
| 18424 | 313 | qed (auto simp add: closed_in_def ty.supp supp_atm abs_supp) | 
| 18246 | 314 | |
| 315 | lemma subtype_implies_fresh: | |
| 316 | fixes X::"tyvrs" | |
| 317 | assumes a1: "\<Gamma> \<turnstile> S <: T" | |
| 318 | and a2: "X\<sharp>\<Gamma>" | |
| 18424 | 319 | shows "X\<sharp>S \<and> X\<sharp>T" | 
| 18246 | 320 | proof - | 
| 321 | from a1 have "\<turnstile> \<Gamma> ok" by (rule subtype_implies_ok) | |
| 18621 | 322 | with a2 have "X\<sharp>domain(\<Gamma>)" by (simp add: fresh_domain) | 
| 18424 | 323 | moreover | 
| 18246 | 324 | from a1 have "S closed_in \<Gamma> \<and> T closed_in \<Gamma>" by (rule subtype_implies_closed) | 
| 18424 | 325 | hence "supp S \<subseteq> ((supp (domain \<Gamma>))::tyvrs set)" | 
| 18621 | 326 | and "supp T \<subseteq> ((supp (domain \<Gamma>))::tyvrs set)" by (simp_all add: domain_supp closed_in_def) | 
| 18424 | 327 | ultimately show "X\<sharp>S \<and> X\<sharp>T" by (force simp add: supp_prod fresh_def) | 
| 18246 | 328 | qed | 
| 329 | ||
| 22730 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22542diff
changeset | 330 | equivariance subtype_of | 
| 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22542diff
changeset | 331 | |
| 22537 | 332 | nominal_inductive subtype_of | 
| 333 | by (simp_all add: abs_fresh subtype_implies_fresh) | |
| 18628 | 334 | |
| 22537 | 335 | thm subtype_of.strong_induct | 
| 18246 | 336 | |
| 18621 | 337 | section {* Reflexivity of Subtyping *}
 | 
| 18246 | 338 | |
| 339 | lemma subtype_reflexivity: | |
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 340 | assumes a: "\<turnstile> \<Gamma> ok" | 
| 18424 | 341 | and b: "T closed_in \<Gamma>" | 
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 342 | shows "\<Gamma> \<turnstile> T <: T" | 
| 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 343 | using a b | 
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26091diff
changeset | 344 | proof(nominal_induct T avoiding: \<Gamma> rule: ty.strong_induct) | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 345 | case (Forall X T\<^isub>1 T\<^isub>2) | 
| 18747 | 346 | have ih_T\<^isub>1: "\<And>\<Gamma>. \<lbrakk>\<turnstile> \<Gamma> ok; T\<^isub>1 closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> T\<^isub>1 <: T\<^isub>1" by fact | 
| 347 | have ih_T\<^isub>2: "\<And>\<Gamma>. \<lbrakk>\<turnstile> \<Gamma> ok; T\<^isub>2 closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> T\<^isub>2 <: T\<^isub>2" by fact | |
| 18424 | 348 | have fresh_cond: "X\<sharp>\<Gamma>" by fact | 
| 18621 | 349 | hence fresh_domain: "X\<sharp>(domain \<Gamma>)" by (simp add: fresh_domain) | 
| 18424 | 350 | have "(\<forall>[X<:T\<^isub>2].T\<^isub>1) closed_in \<Gamma>" by fact | 
| 351 | hence closed\<^isub>T\<^isub>2: "T\<^isub>2 closed_in \<Gamma>" and closed\<^isub>T\<^isub>1: "T\<^isub>1 closed_in ((X,T\<^isub>2)#\<Gamma>)" | |
| 352 | by (auto simp add: closed_in_def ty.supp abs_supp) | |
| 353 | have ok: "\<turnstile> \<Gamma> ok" by fact | |
| 18621 | 354 | hence ok': "\<turnstile> ((X,T\<^isub>2)#\<Gamma>) ok" using closed\<^isub>T\<^isub>2 fresh_domain by simp | 
| 18424 | 355 | have "\<Gamma> \<turnstile> T\<^isub>2 <: T\<^isub>2" using ih_T\<^isub>2 closed\<^isub>T\<^isub>2 ok by simp | 
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 356 | moreover | 
| 18424 | 357 | have "((X,T\<^isub>2)#\<Gamma>) \<turnstile> T\<^isub>1 <: T\<^isub>1" using ih_T\<^isub>1 closed\<^isub>T\<^isub>1 ok' by simp | 
| 358 | ultimately show "\<Gamma> \<turnstile> \<forall>[X<:T\<^isub>2].T\<^isub>1 <: \<forall>[X<:T\<^isub>2].T\<^isub>1" using fresh_cond | |
| 18621 | 359 | by (simp add: subtype_of.S_Forall) | 
| 18246 | 360 | qed (auto simp add: closed_in_def ty.supp supp_atm) | 
| 361 | ||
| 18621 | 362 | lemma subtype_reflexivity_semiautomated: | 
| 18305 
a780f9c1538b
changed everything until the interesting transitivity_narrowing
 urbanc parents: 
18269diff
changeset | 363 | assumes a: "\<turnstile> \<Gamma> ok" | 
| 
a780f9c1538b
changed everything until the interesting transitivity_narrowing
 urbanc parents: 
18269diff
changeset | 364 | and b: "T closed_in \<Gamma>" | 
| 
a780f9c1538b
changed everything until the interesting transitivity_narrowing
 urbanc parents: 
18269diff
changeset | 365 | shows "\<Gamma> \<turnstile> T <: T" | 
| 
a780f9c1538b
changed everything until the interesting transitivity_narrowing
 urbanc parents: 
18269diff
changeset | 366 | using a b | 
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26091diff
changeset | 367 | apply(nominal_induct T avoiding: \<Gamma> rule: ty.strong_induct) | 
| 18747 | 368 | apply(auto simp add: ty.supp abs_supp supp_atm closed_in_def) | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 369 |   --{* Too bad that this instantiation cannot be found automatically by
 | 
| 18621 | 370 |   \isakeyword{auto}; \isakeyword{blast} would find it if we had not used 
 | 
| 18628 | 371 |   an explicit definition for @{text "closed_in_def"}. *}
 | 
| 18305 
a780f9c1538b
changed everything until the interesting transitivity_narrowing
 urbanc parents: 
18269diff
changeset | 372 | apply(drule_tac x="(tyvrs, ty2)#\<Gamma>" in meta_spec) | 
| 18747 | 373 | apply(force dest: fresh_domain simp add: closed_in_def) | 
| 18246 | 374 | done | 
| 375 | ||
| 18747 | 376 | |
| 18628 | 377 | section {* Weakening *}
 | 
| 18246 | 378 | |
| 18628 | 379 | text {* In order to prove weakening we introduce the notion of a type-context extending 
 | 
| 380 | another. This generalization seems to make the proof for weakening to be | |
| 381 | smoother than if we had strictly adhered to the version in the POPLmark-paper. *} | |
| 18246 | 382 | |
| 383 | constdefs | |
| 384 |   extends :: "ty_context \<Rightarrow> ty_context \<Rightarrow> bool" ("_ extends _" [100,100] 100)
 | |
| 385 | "\<Delta> extends \<Gamma> \<equiv> \<forall>X Q. (X,Q)\<in>set \<Gamma> \<longrightarrow> (X,Q)\<in>set \<Delta>" | |
| 386 | ||
| 387 | lemma extends_domain: | |
| 388 | assumes a: "\<Delta> extends \<Gamma>" | |
| 389 | shows "domain \<Gamma> \<subseteq> domain \<Delta>" | |
| 390 | using a | |
| 391 | apply (auto simp add: extends_def) | |
| 392 | apply (drule domain_existence) | |
| 393 | apply (force simp add: domain_inclusion) | |
| 394 | done | |
| 395 | ||
| 396 | lemma extends_closed: | |
| 397 | assumes a1: "T closed_in \<Gamma>" | |
| 398 | and a2: "\<Delta> extends \<Gamma>" | |
| 399 | shows "T closed_in \<Delta>" | |
| 400 | using a1 a2 | |
| 401 | by (auto dest: extends_domain simp add: closed_in_def) | |
| 402 | ||
| 18424 | 403 | lemma extends_memb: | 
| 404 | assumes a: "\<Delta> extends \<Gamma>" | |
| 405 | and b: "(X,T) \<in> set \<Gamma>" | |
| 406 | shows "(X,T) \<in> set \<Delta>" | |
| 407 | using a b by (simp add: extends_def) | |
| 408 | ||
| 18246 | 409 | lemma weakening: | 
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 410 | assumes a: "\<Gamma> \<turnstile> S <: T" | 
| 18424 | 411 | and b: "\<turnstile> \<Delta> ok" | 
| 412 | and c: "\<Delta> extends \<Gamma>" | |
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 413 | shows "\<Delta> \<turnstile> S <: T" | 
| 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 414 | using a b c | 
| 22537 | 415 | proof (nominal_induct \<Gamma> S T avoiding: \<Delta> rule: subtype_of.strong_induct) | 
| 18305 
a780f9c1538b
changed everything until the interesting transitivity_narrowing
 urbanc parents: 
18269diff
changeset | 416 | case (S_Top \<Gamma> S) | 
| 18246 | 417 | have lh_drv_prem: "S closed_in \<Gamma>" by fact | 
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 418 | have "\<turnstile> \<Delta> ok" by fact | 
| 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 419 | moreover | 
| 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 420 | have "\<Delta> extends \<Gamma>" by fact | 
| 18424 | 421 | hence "S closed_in \<Delta>" using lh_drv_prem by (simp only: extends_closed) | 
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 422 | ultimately show "\<Delta> \<turnstile> S <: Top" by force | 
| 18246 | 423 | next | 
| 22537 | 424 | case (S_Var X S \<Gamma> T) | 
| 18246 | 425 | have lh_drv_prem: "(X,S) \<in> set \<Gamma>" by fact | 
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 426 | have ih: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends \<Gamma> \<Longrightarrow> \<Delta> \<turnstile> S <: T" by fact | 
| 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 427 | have ok: "\<turnstile> \<Delta> ok" by fact | 
| 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 428 | have extends: "\<Delta> extends \<Gamma>" by fact | 
| 18424 | 429 | have "(X,S) \<in> set \<Delta>" using lh_drv_prem extends by (simp only: extends_memb) | 
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 430 | moreover | 
| 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 431 | have "\<Delta> \<turnstile> S <: T" using ok extends ih by simp | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 432 | ultimately show "\<Delta> \<turnstile> Tvar X <: T" using ok by force | 
| 18246 | 433 | next | 
| 18305 
a780f9c1538b
changed everything until the interesting transitivity_narrowing
 urbanc parents: 
18269diff
changeset | 434 | case (S_Refl \<Gamma> X) | 
| 18246 | 435 | have lh_drv_prem: "X \<in> domain \<Gamma>" by fact | 
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 436 | have "\<turnstile> \<Delta> ok" by fact | 
| 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 437 | moreover | 
| 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 438 | have "\<Delta> extends \<Gamma>" by fact | 
| 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 439 | hence "X \<in> domain \<Delta>" using lh_drv_prem by (force dest: extends_domain) | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 440 | ultimately show "\<Delta> \<turnstile> Tvar X <: Tvar X" by force | 
| 18246 | 441 | next | 
| 22537 | 442 | case (S_Arrow \<Gamma> T\<^isub>1 S\<^isub>1 S\<^isub>2 T\<^isub>2) thus "\<Delta> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T\<^isub>1 \<rightarrow> T\<^isub>2" by blast | 
| 18246 | 443 | next | 
| 22537 | 444 | case (S_Forall \<Gamma> T\<^isub>1 S\<^isub>1 X S\<^isub>2 T\<^isub>2) | 
| 18424 | 445 | have fresh_cond: "X\<sharp>\<Delta>" by fact | 
| 18621 | 446 | hence fresh_domain: "X\<sharp>(domain \<Delta>)" by (simp add: fresh_domain) | 
| 18424 | 447 | have ih\<^isub>1: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends \<Gamma> \<Longrightarrow> \<Delta> \<turnstile> T\<^isub>1 <: S\<^isub>1" by fact | 
| 448 | have ih\<^isub>2: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends ((X,T\<^isub>1)#\<Gamma>) \<Longrightarrow> \<Delta> \<turnstile> S\<^isub>2 <: T\<^isub>2" by fact | |
| 449 | have lh_drv_prem: "\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1" by fact | |
| 450 | hence closed\<^isub>T\<^isub>1: "T\<^isub>1 closed_in \<Gamma>" by (simp add: subtype_implies_closed) | |
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 451 | have ok: "\<turnstile> \<Delta> ok" by fact | 
| 18424 | 452 | have ext: "\<Delta> extends \<Gamma>" by fact | 
| 453 | have "T\<^isub>1 closed_in \<Delta>" using ext closed\<^isub>T\<^isub>1 by (simp only: extends_closed) | |
| 18621 | 454 | hence "\<turnstile> ((X,T\<^isub>1)#\<Delta>) ok" using fresh_domain ok by force | 
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 455 | moreover | 
| 18424 | 456 | have "((X,T\<^isub>1)#\<Delta>) extends ((X,T\<^isub>1)#\<Gamma>)" using ext by (force simp add: extends_def) | 
| 457 | ultimately have "((X,T\<^isub>1)#\<Delta>) \<turnstile> S\<^isub>2 <: T\<^isub>2" using ih\<^isub>2 by simp | |
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 458 | moreover | 
| 18424 | 459 | have "\<Delta> \<turnstile> T\<^isub>1 <: S\<^isub>1" using ok ext ih\<^isub>1 by simp | 
| 18621 | 460 | ultimately show "\<Delta> \<turnstile> \<forall>[X<:S\<^isub>1].S\<^isub>2 <: \<forall>[X<:T\<^isub>1].T\<^isub>2" using ok by (force intro: S_Forall) | 
| 18246 | 461 | qed | 
| 462 | ||
| 18650 | 463 | text {* In fact all ``non-binding" cases can be solved automatically: *}
 | 
| 18246 | 464 | |
| 18628 | 465 | lemma weakening_more_automated: | 
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 466 | assumes a: "\<Gamma> \<turnstile> S <: T" | 
| 18424 | 467 | and b: "\<turnstile> \<Delta> ok" | 
| 468 | and c: "\<Delta> extends \<Gamma>" | |
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 469 | shows "\<Delta> \<turnstile> S <: T" | 
| 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 470 | using a b c | 
| 22537 | 471 | proof (nominal_induct \<Gamma> S T avoiding: \<Delta> rule: subtype_of.strong_induct) | 
| 472 | case (S_Forall \<Gamma> T\<^isub>1 S\<^isub>1 X S\<^isub>2 T\<^isub>2) | |
| 18424 | 473 | have fresh_cond: "X\<sharp>\<Delta>" by fact | 
| 18621 | 474 | hence fresh_domain: "X\<sharp>(domain \<Delta>)" by (simp add: fresh_domain) | 
| 18424 | 475 | have ih\<^isub>1: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends \<Gamma> \<Longrightarrow> \<Delta> \<turnstile> T\<^isub>1 <: S\<^isub>1" by fact | 
| 476 | have ih\<^isub>2: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends ((X,T\<^isub>1)#\<Gamma>) \<Longrightarrow> \<Delta> \<turnstile> S\<^isub>2 <: T\<^isub>2" by fact | |
| 477 | have lh_drv_prem: "\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1" by fact | |
| 478 | hence closed\<^isub>T\<^isub>1: "T\<^isub>1 closed_in \<Gamma>" by (simp add: subtype_implies_closed) | |
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 479 | have ok: "\<turnstile> \<Delta> ok" by fact | 
| 18424 | 480 | have ext: "\<Delta> extends \<Gamma>" by fact | 
| 481 | have "T\<^isub>1 closed_in \<Delta>" using ext closed\<^isub>T\<^isub>1 by (simp only: extends_closed) | |
| 18621 | 482 | hence "\<turnstile> ((X,T\<^isub>1)#\<Delta>) ok" using fresh_domain ok by force | 
| 18628 | 483 | moreover | 
| 18424 | 484 | have "((X,T\<^isub>1)#\<Delta>) extends ((X,T\<^isub>1)#\<Gamma>)" using ext by (force simp add: extends_def) | 
| 485 | ultimately have "((X,T\<^isub>1)#\<Delta>) \<turnstile> S\<^isub>2 <: T\<^isub>2" using ih\<^isub>2 by simp | |
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 486 | moreover | 
| 18424 | 487 | have "\<Delta> \<turnstile> T\<^isub>1 <: S\<^isub>1" using ok ext ih\<^isub>1 by simp | 
| 18621 | 488 | ultimately show "\<Delta> \<turnstile> \<forall>[X<:S\<^isub>1].S\<^isub>2 <: \<forall>[X<:T\<^isub>1].T\<^isub>2" using ok by (force intro: S_Forall) | 
| 18424 | 489 | qed (blast intro: extends_closed extends_memb dest: extends_domain)+ | 
| 18246 | 490 | |
| 18628 | 491 | section {* Transitivity and Narrowing *}
 | 
| 492 | ||
| 18650 | 493 | text {* Some inversion lemmas that are needed in the transitivity and narrowing proof.*}
 | 
| 494 | ||
| 495 | lemma S_TopE: | |
| 496 | assumes a: "\<Gamma> \<turnstile> Top <: T" | |
| 497 | shows "T = Top" | |
| 498 | using a by (cases, auto) | |
| 499 | ||
| 500 | lemma S_ArrowE_left: | |
| 501 | assumes a: "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T" | |
| 502 | shows "T = Top \<or> (\<exists>T\<^isub>1 T\<^isub>2. T = T\<^isub>1 \<rightarrow> T\<^isub>2 \<and> \<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1 \<and> \<Gamma> \<turnstile> S\<^isub>2 <: T\<^isub>2)" | |
| 503 | using a by (cases, auto simp add: ty.inject) | |
| 504 | ||
| 505 | lemma S_ForallE_left: | |
| 506 | shows "\<lbrakk>\<Gamma> \<turnstile> \<forall>[X<:S\<^isub>1].S\<^isub>2 <: T; X\<sharp>\<Gamma>; X\<sharp>S\<^isub>1\<rbrakk> | |
| 507 | \<Longrightarrow> T = Top \<or> (\<exists>T\<^isub>1 T\<^isub>2. T = \<forall>[X<:T\<^isub>1].T\<^isub>2 \<and> \<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1 \<and> ((X,T\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2)" | |
| 508 | apply(frule subtype_implies_ok) | |
| 23760 | 509 | apply(ind_cases "\<Gamma> \<turnstile> \<forall>[X<:S\<^isub>1].S\<^isub>2 <: T") | 
| 18650 | 510 | apply(auto simp add: ty.inject alpha) | 
| 511 | apply(rule_tac x="[(X,Xa)]\<bullet>T\<^isub>2" in exI) | |
| 512 | apply(rule conjI) | |
| 513 | apply(rule sym) | |
| 514 | apply(rule pt_bij2[OF pt_tyvrs_inst, OF at_tyvrs_inst]) | |
| 515 | apply(rule pt_tyvrs3) | |
| 516 | apply(simp) | |
| 517 | apply(rule at_ds5[OF at_tyvrs_inst]) | |
| 518 | apply(rule conjI) | |
| 519 | apply(simp add: pt_fresh_left[OF pt_tyvrs_inst, OF at_tyvrs_inst] calc_atm) | |
| 520 | apply(drule_tac \<Gamma>="((Xa,T\<^isub>1)#\<Gamma>)" in subtype_implies_closed)+ | |
| 521 | apply(simp add: closed_in_def) | |
| 522 | apply(drule fresh_domain)+ | |
| 523 | apply(simp add: fresh_def) | |
| 524 | apply(subgoal_tac "X \<notin> (insert Xa (domain \<Gamma>))")(*A*) | |
| 525 | apply(force) | |
| 526 | (*A*)apply(simp add: at_fin_set_supp[OF at_tyvrs_inst, OF finite_domain]) | |
| 527 | (* 2nd conjunct *)apply(frule_tac X="X" in subtype_implies_fresh) | |
| 528 | apply(assumption) | |
| 529 | apply(drule_tac X="Xa" in subtype_implies_fresh) | |
| 530 | apply(assumption) | |
| 531 | apply(simp add: fresh_prod) | |
| 22542 | 532 | apply(drule_tac pi="[(X,Xa)]" in subtype_of.eqvt(2)) | 
| 18650 | 533 | apply(simp add: calc_atm) | 
| 534 | apply(simp add: pt_fresh_fresh[OF pt_tyvrs_inst, OF at_tyvrs_inst]) | |
| 535 | done | |
| 536 | ||
| 537 | text {* Next we prove the transitivity and narrowing for the subtyping-relation. 
 | |
| 18621 | 538 | The POPLmark-paper says the following: | 
| 539 | ||
| 18650 | 540 | \begin{quote}
 | 
| 18621 | 541 | \begin{lemma}[Transitivity and Narrowing] \
 | 
| 542 | \begin{enumerate}
 | |
| 543 | \item If @{term "\<Gamma> \<turnstile> S<:Q"} and @{term "\<Gamma> \<turnstile> Q<:T"}, then @{term "\<Gamma> \<turnstile> S<:T"}.
 | |
| 544 | \item If @{text "\<Gamma>,X<:Q,\<Delta> \<turnstile> M<:N"} and @{term "\<Gamma> \<turnstile> P<:Q"} then @{text "\<Gamma>,X<:P,\<Delta> \<turnstile> M<:N"}.
 | |
| 545 | \end{enumerate}
 | |
| 546 | \end{lemma}
 | |
| 547 | ||
| 548 | The two parts are proved simultaneously, by induction on the size | |
| 549 | of @{term "Q"}.  The argument for part (2) assumes that part (1) has 
 | |
| 550 | been established already for the @{term "Q"} in question; part (1) uses 
 | |
| 551 | part (2) only for strictly smaller @{term "Q"}.
 | |
| 18650 | 552 | \end{quote}
 | 
| 18621 | 553 | |
| 554 | For the induction on the size of @{term "Q"}, we use the induction-rule 
 | |
| 555 | @{text "measure_induct_rule"}:
 | |
| 556 | ||
| 557 | \begin{center}
 | |
| 558 | @{thm measure_induct_rule[of "size_ty",no_vars]}
 | |
| 559 | \end{center}
 | |
| 18410 | 560 | |
| 18628 | 561 | That means in order to show a property @{term "P a"} for all @{term "a"}, 
 | 
| 18650 | 562 | the induct-rule requires to prove that for all @{term x} @{term "P x"} holds using the 
 | 
| 18621 | 563 | assumption that for all @{term y} whose size is strictly smaller than 
 | 
| 564 | that of @{term x} the property @{term "P y"} holds. *}
 | |
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 565 | |
| 18621 | 566 | lemma | 
| 567 | shows trans: "\<Gamma>\<turnstile>S<:Q \<Longrightarrow> \<Gamma>\<turnstile>Q<:T \<Longrightarrow> \<Gamma>\<turnstile>S<:T" | |
| 568 | and narrow: "(\<Delta>@[(X,Q)]@\<Gamma>)\<turnstile>M<:N \<Longrightarrow> \<Gamma>\<turnstile>P<:Q \<Longrightarrow> (\<Delta>@[(X,P)]@\<Gamma>)\<turnstile>M<:N" | |
| 20503 | 569 | proof (induct Q arbitrary: \<Gamma> S T \<Delta> X P M N taking: "size_ty" rule: measure_induct_rule) | 
| 18621 | 570 | case (less Q) | 
| 571 |     --{* \begin{minipage}[t]{0.9\textwidth}
 | |
| 572 | First we mention the induction hypotheses of the outer induction for later | |
| 573 |     reference:\end{minipage}*}
 | |
| 574 | have IH_trans: | |
| 575 | "\<And>Q' \<Gamma> S T. \<lbrakk>size_ty Q' < size_ty Q; \<Gamma>\<turnstile>S<:Q'; \<Gamma>\<turnstile>Q'<:T\<rbrakk> \<Longrightarrow> \<Gamma>\<turnstile>S<:T" by fact | |
| 576 | have IH_narrow: | |
| 577 | "\<And>Q' \<Delta> \<Gamma> X M N P. \<lbrakk>size_ty Q' < size_ty Q; (\<Delta>@[(X,Q')]@\<Gamma>)\<turnstile>M<:N; \<Gamma>\<turnstile>P<:Q'\<rbrakk> | |
| 578 | \<Longrightarrow> (\<Delta>@[(X,P)]@\<Gamma>)\<turnstile>M<:N" by fact | |
| 579 |     --{* \begin{minipage}[t]{0.9\textwidth}
 | |
| 580 | We proceed with the transitivity proof as an auxiliary lemma, because it needs | |
| 581 |     to be referenced in the narrowing proof.\end{minipage}*}
 | |
| 582 | have transitivity_aux: | |
| 583 | "\<And>\<Gamma> S T. \<lbrakk>\<Gamma> \<turnstile> S <: Q; \<Gamma> \<turnstile> Q <: T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> S <: T" | |
| 18246 | 584 | proof - | 
| 18424 | 585 | fix \<Gamma>' S' T | 
| 18621 | 586 |     assume "\<Gamma>' \<turnstile> S' <: Q" --{* left-hand derivation *}
 | 
| 587 |       and  "\<Gamma>' \<turnstile> Q <: T"  --{* right-hand derivation *}
 | |
| 588 | thus "\<Gamma>' \<turnstile> S' <: T" | |
| 22537 | 589 | proof (nominal_induct \<Gamma>' S' Q\<equiv>Q rule: subtype_of.strong_induct) | 
| 18424 | 590 | case (S_Top \<Gamma> S) | 
| 18621 | 591 | 	--{* \begin{minipage}[t]{0.9\textwidth}
 | 
| 592 | 	In this case the left-hand derivation is @{term "\<Gamma> \<turnstile> S <: Top"}, giving
 | |
| 593 | 	us @{term "\<turnstile> \<Gamma> ok"} and @{term "S closed_in \<Gamma>"}. This case is straightforward, 
 | |
| 594 | 	because the right-hand derivation must be of the form @{term "\<Gamma> \<turnstile> Top <: Top"} 
 | |
| 595 | 	giving us the equation @{term "T = Top"}.\end{minipage}*}
 | |
| 18424 | 596 | hence rh_drv: "\<Gamma> \<turnstile> Top <: T" by simp | 
| 597 | hence T_inst: "T = Top" by (simp add: S_TopE) | |
| 18621 | 598 | have "\<turnstile> \<Gamma> ok" | 
| 23393 | 599 | and "S closed_in \<Gamma>" by fact+ | 
| 18621 | 600 | hence "\<Gamma> \<turnstile> S <: Top" by (simp add: subtype_of.S_Top) | 
| 18424 | 601 | thus "\<Gamma> \<turnstile> S <: T" using T_inst by simp | 
| 18246 | 602 | next | 
| 22537 | 603 | case (S_Var Y U \<Gamma>) | 
| 18621 | 604 | 	-- {* \begin{minipage}[t]{0.9\textwidth}
 | 
| 605 | 	In this case the left-hand derivation is @{term "\<Gamma> \<turnstile> Tvar Y <: Q"} 
 | |
| 606 | 	with @{term "S = Tvar Y"}. We have therefore @{term "(Y,U)"} 
 | |
| 18650 | 607 | 	is in @{term "\<Gamma>"} and by inner induction hypothesis that @{term "\<Gamma> \<turnstile> U <: T"}. 
 | 
| 18621 | 608 | 	By @{text "S_Var"} follows @{term "\<Gamma> \<turnstile> Tvar Y <: T"}.\end{minipage}*}
 | 
| 18424 | 609 | hence IH_inner: "\<Gamma> \<turnstile> U <: T" by simp | 
| 18621 | 610 | have "(Y,U) \<in> set \<Gamma>" by fact | 
| 611 | with IH_inner show "\<Gamma> \<turnstile> Tvar Y <: T" by (simp add: subtype_of.S_Var) | |
| 18246 | 612 | next | 
| 18424 | 613 | case (S_Refl \<Gamma> X) | 
| 18621 | 614 | 	--{* \begin{minipage}[t]{0.9\textwidth}
 | 
| 615 |         In this case the left-hand derivation is @{term "\<Gamma>\<turnstile>(Tvar X) <: (Tvar X)"} with
 | |
| 616 |         @{term "Q=Tvar X"}. The goal then follows immediately from the right-hand 
 | |
| 617 | 	derivation.\end{minipage}*}
 | |
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 618 | thus "\<Gamma> \<turnstile> Tvar X <: T" by simp | 
| 18246 | 619 | next | 
| 22537 | 620 | case (S_Arrow \<Gamma> Q\<^isub>1 S\<^isub>1 S\<^isub>2 Q\<^isub>2) | 
| 18621 | 621 | 	--{* \begin{minipage}[t]{0.9\textwidth}
 | 
| 622 | 	In this case the left-hand derivation is @{term "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: Q\<^isub>1 \<rightarrow> Q\<^isub>2"} with
 | |
| 623 |         @{term "S\<^isub>1\<rightarrow>S\<^isub>2=S"} and @{term "Q\<^isub>1\<rightarrow>Q\<^isub>2=Q"}. We know that the @{text "size_ty"} of 
 | |
| 624 | 	@{term Q\<^isub>1} and @{term Q\<^isub>2} is smaller than that of @{term Q};
 | |
| 625 | 	so we can apply the outer induction hypotheses for @{term Q\<^isub>1} and @{term Q\<^isub>2}. 
 | |
| 626 | 	We also have the sub-derivations  @{term "\<Gamma>\<turnstile>Q\<^isub>1<:S\<^isub>1"} and @{term "\<Gamma>\<turnstile>S\<^isub>2<:Q\<^isub>2"}.
 | |
| 18628 | 627 | 	The right-hand derivation is @{term "\<Gamma> \<turnstile> Q\<^isub>1 \<rightarrow> Q\<^isub>2 <: T"}. There exist types 
 | 
| 18621 | 628 | 	@{text "T\<^isub>1,T\<^isub>2"} such that @{term "T=Top \<or> T=T\<^isub>1\<rightarrow>T\<^isub>2"}. The @{term "Top"}-case is 
 | 
| 629 | 	straightforward once we know @{term "(S\<^isub>1 \<rightarrow> S\<^isub>2) closed_in \<Gamma>"} and @{term "\<turnstile> \<Gamma> ok"}. 
 | |
| 630 | 	In the other case we have the sub-derivations @{term "\<Gamma>\<turnstile>T\<^isub>1<:Q\<^isub>1"} and @{term "\<Gamma>\<turnstile>Q\<^isub>2<:T\<^isub>2"}. 
 | |
| 631 | 	Using the outer induction hypothesis for transitivity we can derive @{term "\<Gamma>\<turnstile>T\<^isub>1<:S\<^isub>1"} 
 | |
| 632 | 	and @{term "\<Gamma>\<turnstile>S\<^isub>2<:T\<^isub>2"}. By rule @{text "S_Arrow"} follows @{term "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T\<^isub>1 \<rightarrow> T\<^isub>2"},
 | |
| 633 | 	which is @{term "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T\<^isub>"}.\end{minipage}*}
 | |
| 634 | hence rh_drv: "\<Gamma> \<turnstile> Q\<^isub>1 \<rightarrow> Q\<^isub>2 <: T" by simp | |
| 635 | from `Q\<^isub>1 \<rightarrow> Q\<^isub>2 = Q` | |
| 636 | 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 | |
| 637 | have lh_drv_prm\<^isub>1: "\<Gamma> \<turnstile> Q\<^isub>1 <: S\<^isub>1" by fact | |
| 638 | have lh_drv_prm\<^isub>2: "\<Gamma> \<turnstile> S\<^isub>2 <: Q\<^isub>2" by fact | |
| 639 | from rh_drv have "T=Top \<or> (\<exists>T\<^isub>1 T\<^isub>2. T=T\<^isub>1\<rightarrow>T\<^isub>2 \<and> \<Gamma>\<turnstile>T\<^isub>1<:Q\<^isub>1 \<and> \<Gamma>\<turnstile>Q\<^isub>2<:T\<^isub>2)" | |
| 640 | by (simp add: S_ArrowE_left) | |
| 641 | moreover | |
| 642 | have "S\<^isub>1 closed_in \<Gamma>" and "S\<^isub>2 closed_in \<Gamma>" | |
| 643 | using lh_drv_prm\<^isub>1 lh_drv_prm\<^isub>2 by (simp_all add: subtype_implies_closed) | |
| 644 | hence "(S\<^isub>1 \<rightarrow> S\<^isub>2) closed_in \<Gamma>" by (simp add: closed_in_def ty.supp) | |
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 645 | moreover | 
| 18424 | 646 | have "\<turnstile> \<Gamma> ok" using rh_drv by (rule subtype_implies_ok) | 
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 647 | moreover | 
| 18621 | 648 |       { assume "\<exists>T\<^isub>1 T\<^isub>2. T=T\<^isub>1\<rightarrow>T\<^isub>2 \<and> \<Gamma>\<turnstile>T\<^isub>1<:Q\<^isub>1 \<and> \<Gamma>\<turnstile>Q\<^isub>2<:T\<^isub>2"
 | 
| 649 | then obtain T\<^isub>1 T\<^isub>2 | |
| 650 | where T_inst: "T = T\<^isub>1 \<rightarrow> T\<^isub>2" | |
| 651 | and rh_drv_prm\<^isub>1: "\<Gamma> \<turnstile> T\<^isub>1 <: Q\<^isub>1" | |
| 652 | and rh_drv_prm\<^isub>2: "\<Gamma> \<turnstile> Q\<^isub>2 <: T\<^isub>2" by force | |
| 653 | from IH_trans[of "Q\<^isub>1"] | |
| 654 | have "\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1" using Q\<^isub>1\<^isub>2_less rh_drv_prm\<^isub>1 lh_drv_prm\<^isub>1 by simp | |
| 18246 | 655 | moreover | 
| 18621 | 656 | from IH_trans[of "Q\<^isub>2"] | 
| 657 | have "\<Gamma> \<turnstile> 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 | |
| 658 | ultimately have "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T\<^isub>1 \<rightarrow> T\<^isub>2" by (simp add: subtype_of.S_Arrow) | |
| 659 | hence "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T" using T_inst by simp | |
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 660 | } | 
| 18621 | 661 | ultimately show "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T" by blast | 
| 18246 | 662 | next | 
| 22537 | 663 | case (S_Forall \<Gamma> Q\<^isub>1 S\<^isub>1 X S\<^isub>2 Q\<^isub>2) | 
| 18621 | 664 | 	--{* \begin{minipage}[t]{0.9\textwidth}
 | 
| 665 | 	In this case the left-hand derivation is @{text "\<Gamma>\<turnstile>\<forall>[X<:S\<^isub>1].S\<^isub>2 <: \<forall>[X<:Q\<^isub>1].Q\<^isub>2"} with 
 | |
| 666 | 	@{text "\<forall>[X<:S\<^isub>1].S\<^isub>2=S"} and @{text "\<forall>[X<:Q\<^isub>1].Q\<^isub>2=Q"}. We therefore have the sub-derivations  
 | |
| 667 | 	@{term "\<Gamma>\<turnstile>Q\<^isub>1<:S\<^isub>1"} and @{term "((X,Q\<^isub>1)#\<Gamma>)\<turnstile>S\<^isub>2<:Q\<^isub>2"}. Since @{term "X"} is a binder, we
 | |
| 668 | assume that it is sufficiently fresh; in particular we have the freshness conditions | |
| 18650 | 669 | 	@{term "X\<sharp>\<Gamma>"} and @{term "X\<sharp>Q\<^isub>1"} (these assumptions are provided by the strong 
 | 
| 670 | 	induction-rule @{text "subtype_of_induct"}). We know that the @{text "size_ty"} of 
 | |
| 18621 | 671 | 	@{term Q\<^isub>1} and @{term Q\<^isub>2} is smaller than that of @{term Q};
 | 
| 672 | 	so we can apply the outer induction hypotheses for @{term Q\<^isub>1} and @{term Q\<^isub>2}. 
 | |
| 673 | 	The right-hand derivation is @{text "\<Gamma> \<turnstile> \<forall>[X<:Q\<^isub>1].Q\<^isub>2 <: T"}. Since @{term "X\<sharp>\<Gamma>"} 
 | |
| 674 | 	and @{term "X\<sharp>Q\<^isub>1"} there exists types @{text "T\<^isub>1,T\<^isub>2"} such that 
 | |
| 675 | 	@{text "T=Top \<or> T=\<forall>[X<:T\<^isub>1].T\<^isub>2"}. The @{term "Top"}-case is straightforward once we know 
 | |
| 676 | 	@{text "(\<forall>[X<:S\<^isub>1].S\<^isub>2) closed_in \<Gamma>"} and @{term "\<turnstile> \<Gamma> ok"}. In the other case we have 
 | |
| 18628 | 677 | 	the sub-derivations @{term "\<Gamma>\<turnstile>T\<^isub>1<:Q\<^isub>1"} and @{term "((X,T\<^isub>1)#\<Gamma>)\<turnstile>Q\<^isub>2<:T\<^isub>2"}. Using the outer 
 | 
| 18621 | 678 | 	induction hypothesis for transitivity we can derive @{term "\<Gamma>\<turnstile>T\<^isub>1<:S\<^isub>1"}. From the outer 
 | 
| 18628 | 679 | 	induction for narrowing we get @{term "((X,T\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: Q\<^isub>2"} and then using again 
 | 
| 680 | 	induction for transitivity we obtain @{term "((X,T\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2"}. By rule 
 | |
| 681 | 	@{text "S_Forall"} and the freshness condition @{term "X\<sharp>\<Gamma>"} follows 
 | |
| 18650 | 682 | 	@{text "\<Gamma> \<turnstile> \<forall>[X<:S\<^isub>1].S\<^isub>2 <: \<forall>[X<:T\<^isub>1].T\<^isub>2"}, which is @{text "\<Gamma> \<turnstile>  \<forall>[X<:S\<^isub>1].S\<^isub>2 <: T\<^isub>"}.
 | 
| 18628 | 683 | 	\end{minipage}*}
 | 
| 18621 | 684 | hence rh_drv: "\<Gamma> \<turnstile> \<forall>[X<:Q\<^isub>1].Q\<^isub>2 <: T" by simp | 
| 685 | have lh_drv_prm\<^isub>1: "\<Gamma> \<turnstile> Q\<^isub>1 <: S\<^isub>1" by fact | |
| 686 | have lh_drv_prm\<^isub>2: "((X,Q\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: Q\<^isub>2" by fact | |
| 22537 | 687 | have "X\<sharp>\<Gamma>" by fact | 
| 688 | then have fresh_cond: "X\<sharp>\<Gamma>" "X\<sharp>Q\<^isub>1" using lh_drv_prm\<^isub>1 by (simp_all add: subtype_implies_fresh) | |
| 18621 | 689 | from `\<forall>[X<:Q\<^isub>1].Q\<^isub>2 = Q` | 
| 20395 
9a60e3151244
added definition for size and substitution using the recursion
 urbanc parents: 
19972diff
changeset | 690 | 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 | 
| 18621 | 691 | from rh_drv | 
| 692 | have "T=Top \<or> (\<exists>T\<^isub>1 T\<^isub>2. T=\<forall>[X<:T\<^isub>1].T\<^isub>2 \<and> \<Gamma>\<turnstile>T\<^isub>1<:Q\<^isub>1 \<and> ((X,T\<^isub>1)#\<Gamma>)\<turnstile>Q\<^isub>2<:T\<^isub>2)" | |
| 693 | using fresh_cond by (simp add: S_ForallE_left) | |
| 694 | moreover | |
| 695 | have "S\<^isub>1 closed_in \<Gamma>" and "S\<^isub>2 closed_in ((X,Q\<^isub>1)#\<Gamma>)" | |
| 696 | using lh_drv_prm\<^isub>1 lh_drv_prm\<^isub>2 by (simp_all add: subtype_implies_closed) | |
| 697 | hence "(\<forall>[X<:S\<^isub>1].S\<^isub>2) closed_in \<Gamma>" by (force simp add: closed_in_def ty.supp abs_supp) | |
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 698 | moreover | 
| 18424 | 699 | have "\<turnstile> \<Gamma> ok" using rh_drv by (rule subtype_implies_ok) | 
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 700 | moreover | 
| 18621 | 701 |       { assume "\<exists>T\<^isub>1 T\<^isub>2. T=\<forall>[X<:T\<^isub>1].T\<^isub>2 \<and> \<Gamma>\<turnstile>T\<^isub>1<:Q\<^isub>1 \<and> ((X,T\<^isub>1)#\<Gamma>)\<turnstile>Q\<^isub>2<:T\<^isub>2"
 | 
| 702 | then obtain T\<^isub>1 T\<^isub>2 | |
| 703 | where T_inst: "T = \<forall>[X<:T\<^isub>1].T\<^isub>2" | |
| 704 | and rh_drv_prm\<^isub>1: "\<Gamma> \<turnstile> T\<^isub>1 <: Q\<^isub>1" | |
| 705 | and rh_drv_prm\<^isub>2:"((X,T\<^isub>1)#\<Gamma>) \<turnstile> Q\<^isub>2 <: T\<^isub>2" by force | |
| 706 | from IH_trans[of "Q\<^isub>1"] | |
| 707 | have "\<Gamma> \<turnstile> 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 | |
| 708 | moreover | |
| 709 | from IH_narrow[of "Q\<^isub>1" "[]"] | |
| 710 | have "((X,T\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: Q\<^isub>2" using Q\<^isub>1\<^isub>2_less lh_drv_prm\<^isub>2 rh_drv_prm\<^isub>1 by simp | |
| 711 | with IH_trans[of "Q\<^isub>2"] | |
| 712 | have "((X,T\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2" using Q\<^isub>1\<^isub>2_less rh_drv_prm\<^isub>2 by simp | |
| 713 | ultimately have "\<Gamma> \<turnstile> \<forall>[X<:S\<^isub>1].S\<^isub>2 <: \<forall>[X<:T\<^isub>1].T\<^isub>2" | |
| 714 | using fresh_cond by (simp add: subtype_of.S_Forall) | |
| 715 | hence "\<Gamma> \<turnstile> \<forall>[X<:S\<^isub>1].S\<^isub>2 <: T" using T_inst by simp | |
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 716 | } | 
| 18621 | 717 | ultimately show "\<Gamma> \<turnstile> \<forall>[X<:S\<^isub>1].S\<^isub>2 <: T" by blast | 
| 18246 | 718 | qed | 
| 719 | qed | |
| 720 | ||
| 18621 | 721 |   { --{* The transitivity proof is now by the auxiliary lemma. *}
 | 
| 722 | case 1 | |
| 723 | have "\<Gamma> \<turnstile> S <: Q" | |
| 23393 | 724 | and "\<Gamma> \<turnstile> Q <: T" by fact+ | 
| 18621 | 725 | thus "\<Gamma> \<turnstile> S <: T" by (rule transitivity_aux) | 
| 726 | next | |
| 727 |     --{* The narrowing proof proceeds by an induction over @{term "(\<Delta>@[(X,Q)]@\<Gamma>) \<turnstile> M <: N"}. *}
 | |
| 728 | case 2 | |
| 729 |     have  "(\<Delta>@[(X,Q)]@\<Gamma>) \<turnstile> M <: N" --{* left-hand derivation *}
 | |
| 23393 | 730 |       and "\<Gamma> \<turnstile> P<:Q" by fact+ --{* right-hand derivation *}
 | 
| 18621 | 731 | thus "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> M <: N" | 
| 22537 | 732 | proof (nominal_induct \<Gamma>\<equiv>"\<Delta>@[(X,Q)]@\<Gamma>" M N avoiding: \<Delta> \<Gamma> X rule: subtype_of.strong_induct) | 
| 18424 | 733 | case (S_Top _ S \<Delta> \<Gamma> X) | 
| 18621 | 734 | 	--{* \begin{minipage}[t]{0.9\textwidth}
 | 
| 735 | 	In this case the left-hand derivation is @{term "(\<Delta>@[(X,Q)]@\<Gamma>) \<turnstile> S <: Top"}. We show
 | |
| 736 | 	that the context @{term "\<Delta>@[(X,P)]@\<Gamma>"} is ok and that @{term S} is closed in 
 | |
| 737 | 	@{term "\<Delta>@[(X,P)]@\<Gamma>"}. Then we can apply the @{text "S_Top"}-rule.\end{minipage}*}
 | |
| 738 | hence lh_drv_prm\<^isub>1: "\<turnstile> (\<Delta>@[(X,Q)]@\<Gamma>) ok" | |
| 739 | and lh_drv_prm\<^isub>2: "S closed_in (\<Delta>@[(X,Q)]@\<Gamma>)" by simp_all | |
| 18424 | 740 | have rh_drv: "\<Gamma> \<turnstile> P <: Q" by fact | 
| 741 | hence "P closed_in \<Gamma>" by (simp add: subtype_implies_closed) | |
| 18621 | 742 | with lh_drv_prm\<^isub>1 have "\<turnstile> (\<Delta>@[(X,P)]@\<Gamma>) ok" by (simp add: replace_type) | 
| 18412 | 743 | moreover | 
| 18621 | 744 | from lh_drv_prm\<^isub>2 have "S closed_in (\<Delta>@[(X,P)]@\<Gamma>)" | 
| 745 | by (simp add: closed_in_def domain_append) | |
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 746 | ultimately show "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> S <: Top" by (simp add: subtype_of.S_Top) | 
| 18246 | 747 | next | 
| 22537 | 748 | case (S_Var Y S _ N \<Delta> \<Gamma> X) | 
| 18621 | 749 | 	--{* \begin{minipage}[t]{0.9\textwidth}
 | 
| 18628 | 750 | 	In this case the left-hand derivation is @{term "(\<Delta>@[(X,Q)]@\<Gamma>) \<turnstile> Tvar Y <: N"} and
 | 
| 751 | 	by inner induction hypothesis we have @{term "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> S <: N"}. We therefore 
 | |
| 18621 | 752 | 	know that the contexts @{term "\<Delta>@[(X,Q)]@\<Gamma>"} and @{term "\<Delta>@[(X,P)]@\<Gamma>"} are ok, and that 
 | 
| 18628 | 753 | 	@{term "(Y,S)"} is in @{term "\<Delta>@[(X,Q)]@\<Gamma>"}. We need to show that 
 | 
| 754 | 	@{term "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> Tvar Y <: N"}  holds. In case @{term "X\<noteq>Y"} we know that 
 | |
| 755 | 	@{term "(Y,S)"} is in @{term "\<Delta>@[(X,P)]@\<Gamma>"} and can use the inner induction hypothesis 
 | |
| 756 | 	and rule @{text "S_Var"} to conclude. In case @{term "X=Y"} we can infer that 
 | |
| 757 | 	@{term "S=Q"}; moreover we have that  @{term "(\<Delta>@[(X,P)]@\<Gamma>) extends \<Gamma>"} and therefore 
 | |
| 758 | 	by @{text "weakening"} that @{term "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> P <: Q"} holds. By transitivity we 
 | |
| 759 | 	obtain then @{term "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> P <: N"} and can conclude by applying rule 
 | |
| 760 | 	@{text "S_Var"}.\end{minipage}*}
 | |
| 18621 | 761 | hence IH_inner: "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> S <: N" | 
| 762 | and lh_drv_prm: "(Y,S) \<in> set (\<Delta>@[(X,Q)]@\<Gamma>)" | |
| 763 | and rh_drv: "\<Gamma> \<turnstile> P<:Q" | |
| 764 | and ok\<^isub>Q: "\<turnstile> (\<Delta>@[(X,Q)]@\<Gamma>) ok" by (simp_all add: subtype_implies_ok) | |
| 765 | hence ok\<^isub>P: "\<turnstile> (\<Delta>@[(X,P)]@\<Gamma>) ok" by (simp add: subtype_implies_ok) | |
| 766 | show "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> Tvar Y <: N" | |
| 767 | proof (cases "X=Y") | |
| 768 | case False | |
| 769 | have "X\<noteq>Y" by fact | |
| 770 | hence "(Y,S)\<in>set (\<Delta>@[(X,P)]@\<Gamma>)" using lh_drv_prm by simp | |
| 771 | with IH_inner show "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> Tvar Y <: N" by (simp add: subtype_of.S_Var) | |
| 772 | next | |
| 773 | case True | |
| 774 | have memb\<^isub>X\<^isub>Q: "(X,Q)\<in>set (\<Delta>@[(X,Q)]@\<Gamma>)" by simp | |
| 775 | have memb\<^isub>X\<^isub>P: "(X,P)\<in>set (\<Delta>@[(X,P)]@\<Gamma>)" by simp | |
| 776 | have eq: "X=Y" by fact | |
| 777 | hence "S=Q" using ok\<^isub>Q lh_drv_prm memb\<^isub>X\<^isub>Q by (simp only: uniqueness_of_ctxt) | |
| 18424 | 778 | hence "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> Q <: N" using IH_inner by simp | 
| 779 | moreover | |
| 780 | have "(\<Delta>@[(X,P)]@\<Gamma>) extends \<Gamma>" by (simp add: extends_def) | |
| 18621 | 781 | hence "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> P <: Q" using rh_drv ok\<^isub>P by (simp only: weakening) | 
| 782 | ultimately have "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> P <: N" by (simp add: transitivity_aux) | |
| 783 | thus "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> Tvar Y <: N" using memb\<^isub>X\<^isub>P eq by (simp only: subtype_of.S_Var) | |
| 784 | qed | |
| 18246 | 785 | next | 
| 18424 | 786 | case (S_Refl _ Y \<Delta> \<Gamma> X) | 
| 18621 | 787 | 	--{* \begin{minipage}[t]{0.9\textwidth}
 | 
| 788 | 	In this case the left-hand derivation is @{term "(\<Delta>@[(X,Q)]@\<Gamma>) \<turnstile> Tvar Y <: Tvar Y"} and we
 | |
| 18628 | 789 | 	therefore know that @{term "\<Delta>@[(X,Q)]@\<Gamma>"} is ok and that @{term "Y"} is in 
 | 
| 18621 | 790 | 	the domain of @{term "\<Delta>@[(X,Q)]@\<Gamma>"}. We therefore know that @{term "\<Delta>@[(X,P)]@\<Gamma>"} is ok
 | 
| 791 | 	and that @{term Y} is in the domain of @{term "\<Delta>@[(X,P)]@\<Gamma>"}. We can conclude by applying 
 | |
| 792 | 	rule @{text "S_Refl"}.\end{minipage}*}
 | |
| 793 | hence lh_drv_prm\<^isub>1: "\<turnstile> (\<Delta>@[(X,Q)]@\<Gamma>) ok" | |
| 794 | and lh_drv_prm\<^isub>2: "Y \<in> domain (\<Delta>@[(X,Q)]@\<Gamma>)" by simp_all | |
| 18424 | 795 | have "\<Gamma> \<turnstile> P <: Q" by fact | 
| 796 | hence "P closed_in \<Gamma>" by (simp add: subtype_implies_closed) | |
| 18621 | 797 | with lh_drv_prm\<^isub>1 have "\<turnstile> (\<Delta>@[(X,P)]@\<Gamma>) ok" by (simp add: replace_type) | 
| 18424 | 798 | moreover | 
| 18621 | 799 | from lh_drv_prm\<^isub>2 have "Y \<in> domain (\<Delta>@[(X,P)]@\<Gamma>)" by (simp add: domain_append) | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 800 | ultimately show "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> Tvar Y <: Tvar Y" by (simp add: subtype_of.S_Refl) | 
| 18246 | 801 | next | 
| 22537 | 802 | case (S_Arrow _ S\<^isub>1 Q\<^isub>1 Q\<^isub>2 S\<^isub>2 \<Delta> \<Gamma> X) | 
| 18621 | 803 | 	--{* \begin{minipage}[t]{0.9\textwidth}
 | 
| 804 | 	In this case the left-hand derivation is @{term "(\<Delta>@[(X,Q)]@\<Gamma>) \<turnstile> Q\<^isub>1 \<rightarrow> Q\<^isub>2 <: S\<^isub>1 \<rightarrow> S\<^isub>2"} 
 | |
| 805 | 	and the proof is trivial.\end{minipage}*}
 | |
| 806 | thus "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> Q\<^isub>1 \<rightarrow> Q\<^isub>2 <: S\<^isub>1 \<rightarrow> S\<^isub>2" by blast | |
| 18424 | 807 | next | 
| 22537 | 808 | case (S_Forall _ T\<^isub>1 S\<^isub>1 Y S\<^isub>2 T\<^isub>2 \<Delta> \<Gamma> X) | 
| 18621 | 809 | 	--{* \begin{minipage}[t]{0.9\textwidth}
 | 
| 18628 | 810 | 	In this case the left-hand derivation is @{text "(\<Delta>@[(X,Q)]@\<Gamma>) \<turnstile> \<forall>[Y<:S\<^isub>1].S\<^isub>2 <: \<forall>[Y<:T\<^isub>1].T\<^isub>2"}
 | 
| 18621 | 811 | 	and therfore we know that the binder @{term Y} is fresh for @{term "\<Delta>@[(X,Q)]@\<Gamma>"}. By
 | 
| 812 | 	the inner induction hypothesis we have that @{term "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> T\<^isub>1 <: S\<^isub>1"} and 
 | |
| 813 | 	@{term "((Y,T\<^isub>1)#\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2"}. Since @{term P} is a subtype of @{term Q}
 | |
| 814 | 	we can infer that @{term Y} is fresh for @{term P} and thus also fresh for 
 | |
| 815 | 	@{term "\<Delta>@[(X,P)]@\<Gamma>"}. We can then conclude by applying rule @{text "S_Forall"}.
 | |
| 816 | 	\end{minipage}*}
 | |
| 817 | hence IH_inner\<^isub>1: "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> T\<^isub>1 <: S\<^isub>1" | |
| 818 | and IH_inner\<^isub>2: "((Y,T\<^isub>1)#\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2" | |
| 819 | and lh_drv_prm: "Y\<sharp>(\<Delta>@[(X,Q)]@\<Gamma>)" by force+ | |
| 18424 | 820 | have rh_drv: "\<Gamma> \<turnstile> P <: Q" by fact | 
| 18621 | 821 | hence "Y\<sharp>P" using lh_drv_prm by (simp only: fresh_list_append subtype_implies_fresh) | 
| 822 | hence "Y\<sharp>(\<Delta>@[(X,P)]@\<Gamma>)" using lh_drv_prm | |
| 18424 | 823 | by (simp add: fresh_list_append fresh_list_cons fresh_prod) | 
| 18621 | 824 | with IH_inner\<^isub>1 IH_inner\<^isub>2 | 
| 825 | show "(\<Delta>@[(X,P)]@\<Gamma>) \<turnstile> \<forall>[Y<:S\<^isub>1].S\<^isub>2 <: \<forall>[Y<:T\<^isub>1].T\<^isub>2" by (simp add: subtype_of.S_Forall) | |
| 18246 | 826 | qed | 
| 18621 | 827 | } | 
| 18246 | 828 | qed | 
| 829 | ||
| 18416 | 830 | end |