| author | wenzelm | 
| Sun, 22 Aug 2010 20:25:15 +0200 | |
| changeset 38582 | 3a6ce43d99b1 | 
| parent 35438 | 5f1ea533158c | 
| child 39246 | 9e58f0499f57 | 
| permissions | -rw-r--r-- | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 1 | (*<*) | 
| 19501 | 2 | theory Fsub | 
| 3 | imports "../Nominal" | |
| 18269 | 4 | begin | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 5 | (*>*) | 
| 18269 | 6 | |
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 7 | text{* Authors: Christian Urban,
 | 
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 8 | Benjamin Pierce, | 
| 18650 | 9 | Dimitrios Vytiniotis | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 10 | Stephanie Weirich | 
| 18650 | 11 | Steve Zdancewic | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 12 | Julien Narboux | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 13 | Stefan Berghofer | 
| 18266 
55c201fe4c95
added an authors section (please let me know if somebody is left out or unhappy)
 urbanc parents: 
18263diff
changeset | 14 | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 15 | with great help from Markus Wenzel. *} | 
| 18246 | 16 | |
| 18621 | 17 | section {* Types for Names, Nominal Datatype Declaration for Types and Terms *}
 | 
| 18424 | 18 | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 19 | no_syntax | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 20 |   "_Map" :: "maplets => 'a ~=> 'b"  ("(1[_])")
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 21 | |
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 22 | text {* The main point of this solution is to use names everywhere (be they bound, 
 | 
| 18621 | 23 |   binding or free). In System \FSUB{} there are two kinds of names corresponding to 
 | 
| 24 | 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 | 25 |   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 | 26 | |
| 18246 | 27 | atom_decl tyvrs vrs | 
| 28 | ||
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 29 | 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 | 30 |   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 | 31 | |
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 32 | 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 | 33 | 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 | 34 |   @{text "\<guillemotleft>\<dots>\<guillemotright>\<dots>"} to indicate where abstractions occur. *}
 | 
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 35 | |
| 18424 | 36 | nominal_datatype ty = | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 37 | Tvar "tyvrs" | 
| 18424 | 38 | | Top | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 39 | | Arrow "ty" "ty" (infixr "\<rightarrow>" 200) | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 40 | | Forall "\<guillemotleft>tyvrs\<guillemotright>ty" "ty" | 
| 18246 | 41 | |
| 18424 | 42 | nominal_datatype trm = | 
| 43 | Var "vrs" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 44 | | Abs "\<guillemotleft>vrs\<guillemotright>trm" "ty" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 45 | | TAbs "\<guillemotleft>tyvrs\<guillemotright>trm" "ty" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 46 | | App "trm" "trm" (infixl "\<cdot>" 200) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 47 | | TApp "trm" "ty" (infixl "\<cdot>\<^sub>\<tau>" 200) | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 48 | |
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 49 | text {* To be polite to the eye, some more familiar notation is introduced. 
 | 
| 18621 | 50 | 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 | 51 | translation rules, instead of syntax annotations at the term-constructors | 
| 18650 | 52 |   as given above for @{term "Arrow"}. *}
 | 
| 18246 | 53 | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 54 | abbreviation | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 55 |   Forall_syn :: "tyvrs \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> ty"  ("(3\<forall>_<:_./ _)" [0, 0, 10] 10) 
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 56 | where | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 57 | "\<forall>X<:T\<^isub>1. T\<^isub>2 \<equiv> ty.Forall X T\<^isub>2 T\<^isub>1" | 
| 18424 | 58 | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 59 | abbreviation | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 60 |   Abs_syn    :: "vrs \<Rightarrow> ty \<Rightarrow> trm \<Rightarrow> trm"  ("(3\<lambda>_:_./ _)" [0, 0, 10] 10) 
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 61 | where | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 62 | "\<lambda>x:T. t \<equiv> trm.Abs x t T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 63 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 64 | abbreviation | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 65 |   TAbs_syn   :: "tyvrs \<Rightarrow> ty \<Rightarrow> trm \<Rightarrow> trm" ("(3\<lambda>_<:_./ _)" [0, 0, 10] 10) 
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 66 | where | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 67 | "\<lambda>X<:T. t \<equiv> trm.TAbs X t T" | 
| 18246 | 68 | |
| 18621 | 69 | text {* Again there are numerous facts that are proved automatically for @{typ "ty"} 
 | 
| 18650 | 70 |   and @{typ "trm"}: for example that the set of free variables, i.e.~the @{text "support"}, 
 | 
| 71 |   is finite. However note that nominal-datatype declarations do \emph{not} define 
 | |
| 72 | ``classical" constructor-based datatypes, but rather define $\alpha$-equivalence | |
| 18621 | 73 |   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 | 74 |   and @{typ "trm"}s are equal: *}
 | 
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 75 | |
| 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 76 | lemma alpha_illustration: | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 77 | shows "(\<forall>X<:T. Tvar X) = (\<forall>Y<:T. Tvar Y)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 78 | and "(\<lambda>x:T. Var x) = (\<lambda>y:T. Var y)" | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 79 | 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 | 80 | |
| 18621 | 81 | section {* SubTyping Contexts *}
 | 
| 18246 | 82 | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 83 | nominal_datatype binding = | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 84 | VarB vrs ty | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 85 | | TVarB tyvrs ty | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 86 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 87 | types env = "binding list" | 
| 18246 | 88 | |
| 18650 | 89 | text {* Typing contexts are represented as lists that ``grow" on the left; we
 | 
| 18621 | 90 | thereby deviating from the convention in the POPLmark-paper. The lists contain | 
| 18650 | 91 | 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 | 92 | |
| 18628 | 93 | text {* In order to state validity-conditions for typing-contexts, the notion of
 | 
| 32011 | 94 |   a @{text "dom"} of a typing-context is handy. *}
 | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 95 | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 96 | nominal_primrec | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 97 | "tyvrs_of" :: "binding \<Rightarrow> tyvrs set" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 98 | where | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 99 |   "tyvrs_of (VarB  x y) = {}"
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 100 | | "tyvrs_of (TVarB x y) = {x}"
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 101 | by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 102 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 103 | nominal_primrec | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 104 | "vrs_of" :: "binding \<Rightarrow> vrs set" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 105 | where | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 106 |   "vrs_of (VarB  x y) = {x}"
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 107 | | "vrs_of (TVarB x y) = {}"
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 108 | by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 109 | |
| 18246 | 110 | consts | 
| 32011 | 111 | "ty_dom" :: "env \<Rightarrow> tyvrs set" | 
| 18246 | 112 | primrec | 
| 32011 | 113 |   "ty_dom [] = {}"
 | 
| 114 | "ty_dom (X#\<Gamma>) = (tyvrs_of X)\<union>(ty_dom \<Gamma>)" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 115 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 116 | consts | 
| 32011 | 117 | "trm_dom" :: "env \<Rightarrow> vrs set" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 118 | primrec | 
| 32011 | 119 |   "trm_dom [] = {}"
 | 
| 120 | "trm_dom (X#\<Gamma>) = (vrs_of X)\<union>(trm_dom \<Gamma>)" | |
| 18246 | 121 | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 122 | lemma vrs_of_eqvt[eqvt]: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 123 | fixes pi ::"tyvrs prm" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 124 | and pi'::"vrs prm" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 125 | shows "pi \<bullet>(tyvrs_of x) = tyvrs_of (pi\<bullet>x)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 126 | and "pi'\<bullet>(tyvrs_of x) = tyvrs_of (pi'\<bullet>x)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 127 | and "pi \<bullet>(vrs_of x) = vrs_of (pi\<bullet>x)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 128 | and "pi'\<bullet>(vrs_of x) = vrs_of (pi'\<bullet>x)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 129 | by (nominal_induct x rule: binding.strong_induct) (simp_all add: tyvrs_of.simps eqvts) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 130 | |
| 32011 | 131 | lemma doms_eqvt[eqvt]: | 
| 18246 | 132 | fixes pi::"tyvrs prm" | 
| 22537 | 133 | and pi'::"vrs prm" | 
| 32011 | 134 | shows "pi \<bullet>(ty_dom \<Gamma>) = ty_dom (pi\<bullet>\<Gamma>)" | 
| 135 | and "pi'\<bullet>(ty_dom \<Gamma>) = ty_dom (pi'\<bullet>\<Gamma>)" | |
| 136 | and "pi \<bullet>(trm_dom \<Gamma>) = trm_dom (pi\<bullet>\<Gamma>)" | |
| 137 | and "pi'\<bullet>(trm_dom \<Gamma>) = trm_dom (pi'\<bullet>\<Gamma>)" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 138 | by (induct \<Gamma>) (simp_all add: eqvts) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 139 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 140 | lemma finite_vrs: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 141 | shows "finite (tyvrs_of x)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 142 | and "finite (vrs_of x)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 143 | by (nominal_induct rule:binding.strong_induct, auto) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 144 | |
| 32011 | 145 | lemma finite_doms: | 
| 146 | shows "finite (ty_dom \<Gamma>)" | |
| 147 | and "finite (trm_dom \<Gamma>)" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 148 | by (induct \<Gamma>, auto simp add: finite_vrs) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 149 | |
| 32011 | 150 | lemma ty_dom_supp: | 
| 151 | shows "(supp (ty_dom \<Gamma>)) = (ty_dom \<Gamma>)" | |
| 152 | and "(supp (trm_dom \<Gamma>)) = (trm_dom \<Gamma>)" | |
| 153 | by (simp only: at_fin_set_supp at_tyvrs_inst at_vrs_inst finite_doms)+ | |
| 18246 | 154 | |
| 32011 | 155 | lemma ty_dom_inclusion: | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 156 | assumes a: "(TVarB X T)\<in>set \<Gamma>" | 
| 32011 | 157 | shows "X\<in>(ty_dom \<Gamma>)" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 158 | using a by (induct \<Gamma>, auto) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 159 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 160 | lemma ty_binding_existence: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 161 | assumes "X \<in> (tyvrs_of a)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 162 | shows "\<exists>T.(TVarB X T=a)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 163 | using assms | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 164 | by (nominal_induct a rule: binding.strong_induct, auto) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 165 | |
| 32011 | 166 | lemma ty_dom_existence: | 
| 167 | assumes a: "X\<in>(ty_dom \<Gamma>)" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 168 | shows "\<exists>T.(TVarB X T)\<in>set \<Gamma>" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 169 | using a | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 170 | apply (induct \<Gamma>, auto) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 171 | apply (subgoal_tac "\<exists>T.(TVarB X T=a)") | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 172 | apply (auto) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 173 | apply (auto simp add: ty_binding_existence) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 174 | done | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 175 | |
| 32011 | 176 | lemma doms_append: | 
| 177 | shows "ty_dom (\<Gamma>@\<Delta>) = ((ty_dom \<Gamma>) \<union> (ty_dom \<Delta>))" | |
| 178 | and "trm_dom (\<Gamma>@\<Delta>) = ((trm_dom \<Gamma>) \<union> (trm_dom \<Delta>))" | |
| 18246 | 179 | by (induct \<Gamma>, auto) | 
| 180 | ||
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 181 | lemma ty_vrs_prm_simp: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 182 | fixes pi::"vrs prm" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 183 | and S::"ty" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 184 | shows "pi\<bullet>S = S" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 185 | by (induct S rule: ty.induct) (auto simp add: calc_atm) | 
| 18246 | 186 | |
| 32011 | 187 | lemma fresh_ty_dom_cons: | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 188 | fixes X::"tyvrs" | 
| 32011 | 189 | shows "X\<sharp>(ty_dom (Y#\<Gamma>)) = (X\<sharp>(tyvrs_of Y) \<and> X\<sharp>(ty_dom \<Gamma>))" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 190 | apply (nominal_induct rule:binding.strong_induct) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 191 | apply (auto) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 192 | apply (simp add: fresh_def supp_def eqvts) | 
| 32011 | 193 | apply (simp add: fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_doms) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 194 | apply (simp add: fresh_def supp_def eqvts) | 
| 32011 | 195 | apply (simp add: fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_doms)+ | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 196 | done | 
| 18246 | 197 | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 198 | lemma tyvrs_fresh: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 199 | fixes X::"tyvrs" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 200 | assumes "X \<sharp> a" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 201 | shows "X \<sharp> tyvrs_of a" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 202 | and "X \<sharp> vrs_of a" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 203 | using assms | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 204 | apply (nominal_induct a rule:binding.strong_induct) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 205 | apply (auto) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 206 | apply (fresh_guess)+ | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 207 | done | 
| 18621 | 208 | |
| 32011 | 209 | lemma fresh_dom: | 
| 18621 | 210 | fixes X::"tyvrs" | 
| 211 | assumes a: "X\<sharp>\<Gamma>" | |
| 32011 | 212 | shows "X\<sharp>(ty_dom \<Gamma>)" | 
| 18621 | 213 | using a | 
| 214 | apply(induct \<Gamma>) | |
| 215 | apply(simp add: fresh_set_empty) | |
| 32011 | 216 | apply(simp only: fresh_ty_dom_cons) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 217 | apply(auto simp add: fresh_prod fresh_list_cons tyvrs_fresh) | 
| 18621 | 218 | done | 
| 219 | ||
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 220 | text {* Not all lists of type @{typ "env"} are well-formed. One condition
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 221 |   requires that in @{term "TVarB X S#\<Gamma>"} all free variables of @{term "S"} must be 
 | 
| 32011 | 222 |   in the @{term "ty_dom"} of @{term "\<Gamma>"}, that is @{term "S"} must be @{text "closed"} 
 | 
| 18650 | 223 |   in @{term "\<Gamma>"}. The set of free variables of @{term "S"} is the 
 | 
| 18621 | 224 |   @{text "support"} of @{term "S"}. *}
 | 
| 18246 | 225 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
34915diff
changeset | 226 | definition "closed_in" :: "ty \<Rightarrow> env \<Rightarrow> bool" ("_ closed'_in _" [100,100] 100) where
 | 
| 32011 | 227 | "S closed_in \<Gamma> \<equiv> (supp S)\<subseteq>(ty_dom \<Gamma>)" | 
| 18246 | 228 | |
| 22537 | 229 | lemma closed_in_eqvt[eqvt]: | 
| 18246 | 230 | fixes pi::"tyvrs prm" | 
| 231 | assumes a: "S closed_in \<Gamma>" | |
| 232 | shows "(pi\<bullet>S) closed_in (pi\<bullet>\<Gamma>)" | |
| 233 | using a | |
| 26091 | 234 | proof - | 
| 235 | from a have "pi\<bullet>(S closed_in \<Gamma>)" by (simp add: perm_bool) | |
| 236 | then show "(pi\<bullet>S) closed_in (pi\<bullet>\<Gamma>)" by (simp add: closed_in_def eqvts) | |
| 18246 | 237 | qed | 
| 238 | ||
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 239 | lemma tyvrs_vrs_prm_simp: | 
| 22537 | 240 | fixes pi::"vrs prm" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 241 | shows "tyvrs_of (pi\<bullet>a) = tyvrs_of a" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 242 | apply (nominal_induct rule:binding.strong_induct) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 243 | apply (simp_all add: eqvts) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 244 | apply (simp add: dj_perm_forget[OF dj_tyvrs_vrs]) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 245 | done | 
| 22537 | 246 | |
| 30986 
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
 Christian Urban <urbanc@in.tum.de> parents: 
30091diff
changeset | 247 | lemma ty_vrs_fresh: | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 248 | fixes x::"vrs" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 249 | and T::"ty" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 250 | shows "x \<sharp> T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 251 | by (simp add: fresh_def supp_def ty_vrs_prm_simp) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 252 | |
| 32011 | 253 | lemma ty_dom_vrs_prm_simp: | 
| 22537 | 254 | fixes pi::"vrs prm" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 255 | and \<Gamma>::"env" | 
| 32011 | 256 | shows "(ty_dom (pi\<bullet>\<Gamma>)) = (ty_dom \<Gamma>)" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 257 | apply(induct \<Gamma>) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 258 | apply (simp add: eqvts) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 259 | apply(simp add: tyvrs_vrs_prm_simp) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 260 | done | 
| 22537 | 261 | |
| 262 | lemma closed_in_eqvt'[eqvt]: | |
| 263 | fixes pi::"vrs prm" | |
| 264 | assumes a: "S closed_in \<Gamma>" | |
| 265 | shows "(pi\<bullet>S) closed_in (pi\<bullet>\<Gamma>)" | |
| 266 | using a | |
| 32011 | 267 | by (simp add: closed_in_def ty_dom_vrs_prm_simp ty_vrs_prm_simp) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 268 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 269 | lemma fresh_vrs_of: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 270 | fixes x::"vrs" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 271 | shows "x\<sharp>vrs_of b = x\<sharp>b" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 272 | by (nominal_induct b rule: binding.strong_induct) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 273 | (simp_all add: fresh_singleton [OF pt_vrs_inst at_vrs_inst] fresh_set_empty ty_vrs_fresh fresh_atm) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 274 | |
| 32011 | 275 | lemma fresh_trm_dom: | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 276 | fixes x::"vrs" | 
| 32011 | 277 | shows "x\<sharp> trm_dom \<Gamma> = x\<sharp>\<Gamma>" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 278 | by (induct \<Gamma>) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 279 | (simp_all add: fresh_set_empty fresh_list_cons | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 280 | fresh_fin_union [OF pt_vrs_inst at_vrs_inst fs_vrs_inst] | 
| 32011 | 281 | finite_doms finite_vrs fresh_vrs_of fresh_list_nil) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 282 | |
| 32011 | 283 | lemma closed_in_fresh: "(X::tyvrs) \<sharp> ty_dom \<Gamma> \<Longrightarrow> T closed_in \<Gamma> \<Longrightarrow> X \<sharp> T" | 
| 284 | by (auto simp add: closed_in_def fresh_def ty_dom_supp) | |
| 22537 | 285 | |
| 18621 | 286 | text {* Now validity of a context is a straightforward inductive definition. *}
 | 
| 287 | ||
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 288 | inductive | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 289 |   valid_rel :: "env \<Rightarrow> bool" ("\<turnstile> _ ok" [100] 100)
 | 
| 22436 | 290 | where | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 291 | valid_nil[simp]: "\<turnstile> [] ok" | 
| 32011 | 292 | | valid_consT[simp]: "\<lbrakk>\<turnstile> \<Gamma> ok; X\<sharp>(ty_dom \<Gamma>); T closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<turnstile> (TVarB X T#\<Gamma>) ok" | 
| 293 | | valid_cons [simp]: "\<lbrakk>\<turnstile> \<Gamma> ok; x\<sharp>(trm_dom \<Gamma>); T closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<turnstile> (VarB x T#\<Gamma>) ok" | |
| 18246 | 294 | |
| 22537 | 295 | equivariance valid_rel | 
| 18246 | 296 | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 297 | declare binding.inject [simp add] | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 298 | declare trm.inject [simp add] | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 299 | |
| 32011 | 300 | inductive_cases validE[elim]: | 
| 301 | "\<turnstile> (TVarB X T#\<Gamma>) ok" | |
| 302 | "\<turnstile> (VarB x T#\<Gamma>) ok" | |
| 303 | "\<turnstile> (b#\<Gamma>) ok" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 304 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 305 | declare binding.inject [simp del] | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 306 | declare trm.inject [simp del] | 
| 18246 | 307 | |
| 18424 | 308 | lemma validE_append: | 
| 309 | assumes a: "\<turnstile> (\<Delta>@\<Gamma>) ok" | |
| 310 | shows "\<turnstile> \<Gamma> ok" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 311 | using a | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 312 | proof (induct \<Delta>) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 313 | case (Cons a \<Gamma>') | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 314 | then show ?case | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 315 | by (nominal_induct a rule:binding.strong_induct) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 316 | (auto elim: validE) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 317 | qed (auto) | 
| 18246 | 318 | |
| 18424 | 319 | lemma replace_type: | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 320 | assumes a: "\<turnstile> (\<Delta>@(TVarB X T)#\<Gamma>) ok" | 
| 18424 | 321 | and b: "S closed_in \<Gamma>" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 322 | shows "\<turnstile> (\<Delta>@(TVarB X S)#\<Gamma>) ok" | 
| 18621 | 323 | using a b | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 324 | proof(induct \<Delta>) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 325 | case Nil | 
| 32011 | 326 | then show ?case by (auto elim: validE intro: valid_cons simp add: doms_append closed_in_def) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 327 | next | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 328 | case (Cons a \<Gamma>') | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 329 | then show ?case | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 330 | by (nominal_induct a rule:binding.strong_induct) | 
| 32011 | 331 | (auto elim: validE intro!: valid_cons simp add: doms_append closed_in_def) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 332 | qed | 
| 18246 | 333 | |
| 18650 | 334 | text {* Well-formed contexts have a unique type-binding for a type-variable. *} 
 | 
| 335 | ||
| 18246 | 336 | lemma uniqueness_of_ctxt: | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 337 | fixes \<Gamma>::"env" | 
| 18412 | 338 | assumes a: "\<turnstile> \<Gamma> ok" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 339 | and b: "(TVarB X T)\<in>set \<Gamma>" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 340 | and c: "(TVarB X S)\<in>set \<Gamma>" | 
| 18412 | 341 | shows "T=S" | 
| 18621 | 342 | using a b c | 
| 343 | proof (induct) | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 344 | case (valid_consT \<Gamma> X' T') | 
| 18621 | 345 | moreover | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 346 |   { fix \<Gamma>'::"env"
 | 
| 32011 | 347 | assume a: "X'\<sharp>(ty_dom \<Gamma>')" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 348 | have "\<not>(\<exists>T.(TVarB X' T)\<in>(set \<Gamma>'))" using a | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 349 | proof (induct \<Gamma>') | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 350 | case (Cons Y \<Gamma>') | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 351 | thus "\<not> (\<exists>T.(TVarB X' T)\<in>set(Y#\<Gamma>'))" | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 352 | by (simp add: fresh_ty_dom_cons | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 353 | fresh_fin_union[OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] | 
| 32011 | 354 | finite_vrs finite_doms, | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 355 | auto simp add: fresh_atm fresh_singleton [OF pt_tyvrs_inst at_tyvrs_inst]) | 
| 18621 | 356 | qed (simp) | 
| 357 | } | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 358 | ultimately show "T=S" by (auto simp add: binding.inject) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 359 | qed (auto) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 360 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 361 | lemma uniqueness_of_ctxt': | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 362 | fixes \<Gamma>::"env" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 363 | assumes a: "\<turnstile> \<Gamma> ok" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 364 | and b: "(VarB x T)\<in>set \<Gamma>" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 365 | and c: "(VarB x S)\<in>set \<Gamma>" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 366 | shows "T=S" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 367 | using a b c | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 368 | proof (induct) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 369 | case (valid_cons \<Gamma> x' T') | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 370 | moreover | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 371 |   { fix \<Gamma>'::"env"
 | 
| 32011 | 372 | assume a: "x'\<sharp>(trm_dom \<Gamma>')" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 373 | have "\<not>(\<exists>T.(VarB x' T)\<in>(set \<Gamma>'))" using a | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 374 | proof (induct \<Gamma>') | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 375 | case (Cons y \<Gamma>') | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 376 | thus "\<not> (\<exists>T.(VarB x' T)\<in>set(y#\<Gamma>'))" | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 377 | by (simp add: fresh_fin_union[OF pt_vrs_inst at_vrs_inst fs_vrs_inst] | 
| 32011 | 378 | finite_vrs finite_doms, | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 379 | auto simp add: fresh_atm fresh_singleton [OF pt_vrs_inst at_vrs_inst]) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 380 | qed (simp) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 381 | } | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 382 | ultimately show "T=S" by (auto simp add: binding.inject) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 383 | qed (auto) | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 384 | |
| 18628 | 385 | section {* Size and Capture-Avoiding Substitution for Types *}
 | 
| 18621 | 386 | |
| 21554 | 387 | nominal_primrec | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 388 | size_ty :: "ty \<Rightarrow> nat" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 389 | where | 
| 21554 | 390 | "size_ty (Tvar X) = 1" | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 391 | | "size_ty (Top) = 1" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 392 | | "size_ty (T1 \<rightarrow> T2) = (size_ty T1) + (size_ty T2) + 1" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 393 | | "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 | 394 | apply (finite_guess)+ | 
| 21554 | 395 | apply (rule TrueI)+ | 
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
21554diff
changeset | 396 | apply (simp add: fresh_nat) | 
| 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
21554diff
changeset | 397 | apply (fresh_guess)+ | 
| 21554 | 398 | done | 
| 20395 
9a60e3151244
added definition for size and substitution using the recursion
 urbanc parents: 
19972diff
changeset | 399 | |
| 21554 | 400 | nominal_primrec | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 401 |   subst_ty :: "ty \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> ty" ("_[_ \<mapsto> _]\<^sub>\<tau>" [300, 0, 0] 300)
 | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
26966diff
changeset | 402 | where | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 403 | "(Tvar X)[Y \<mapsto> T]\<^sub>\<tau> = (if X=Y then T else Tvar X)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 404 | | "(Top)[Y \<mapsto> T]\<^sub>\<tau> = Top" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 405 | | "(T\<^isub>1 \<rightarrow> T\<^isub>2)[Y \<mapsto> T]\<^sub>\<tau> = T\<^isub>1[Y \<mapsto> T]\<^sub>\<tau> \<rightarrow> T\<^isub>2[Y \<mapsto> T]\<^sub>\<tau>" | 
| 32011 | 406 | | "X\<sharp>(Y,T,T\<^isub>1) \<Longrightarrow> (\<forall>X<:T\<^isub>1. T\<^isub>2)[Y \<mapsto> T]\<^sub>\<tau> = (\<forall>X<:T\<^isub>1[Y \<mapsto> T]\<^sub>\<tau>. T\<^isub>2[Y \<mapsto> T]\<^sub>\<tau>)" | 
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
21554diff
changeset | 407 | apply (finite_guess)+ | 
| 21554 | 408 | apply (rule TrueI)+ | 
| 409 | apply (simp add: abs_fresh) | |
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
21554diff
changeset | 410 | apply (fresh_guess)+ | 
| 21554 | 411 | done | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 412 | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 413 | lemma subst_eqvt[eqvt]: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 414 | fixes pi::"tyvrs prm" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 415 | and T::"ty" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 416 | shows "pi\<bullet>(T[X \<mapsto> T']\<^sub>\<tau>) = (pi\<bullet>T)[(pi\<bullet>X) \<mapsto> (pi\<bullet>T')]\<^sub>\<tau>" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 417 | by (nominal_induct T avoiding: X T' rule: ty.strong_induct) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 418 | (perm_simp add: fresh_bij)+ | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 419 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 420 | lemma subst_eqvt'[eqvt]: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 421 | fixes pi::"vrs prm" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 422 | and T::"ty" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 423 | shows "pi\<bullet>(T[X \<mapsto> T']\<^sub>\<tau>) = (pi\<bullet>T)[(pi\<bullet>X) \<mapsto> (pi\<bullet>T')]\<^sub>\<tau>" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 424 | by (nominal_induct T avoiding: X T' rule: ty.strong_induct) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 425 | (perm_simp add: fresh_left)+ | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 426 | |
| 30986 
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
 Christian Urban <urbanc@in.tum.de> parents: 
30091diff
changeset | 427 | lemma type_subst_fresh: | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 428 | fixes X::"tyvrs" | 
| 32011 | 429 | assumes "X\<sharp>T" and "X\<sharp>P" | 
| 430 | shows "X\<sharp>T[Y \<mapsto> P]\<^sub>\<tau>" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 431 | using assms | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 432 | by (nominal_induct T avoiding: X Y P rule:ty.strong_induct) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 433 | (auto simp add: abs_fresh) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 434 | |
| 30986 
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
 Christian Urban <urbanc@in.tum.de> parents: 
30091diff
changeset | 435 | lemma fresh_type_subst_fresh: | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 436 | assumes "X\<sharp>T'" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 437 | shows "X\<sharp>T[X \<mapsto> T']\<^sub>\<tau>" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 438 | using assms | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 439 | by (nominal_induct T avoiding: X T' rule: ty.strong_induct) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 440 | (auto simp add: fresh_atm abs_fresh fresh_nat) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 441 | |
| 32011 | 442 | lemma type_subst_identity: | 
| 443 | "X\<sharp>T \<Longrightarrow> T[X \<mapsto> U]\<^sub>\<tau> = T" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 444 | by (nominal_induct T avoiding: X U rule: ty.strong_induct) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 445 | (simp_all add: fresh_atm abs_fresh) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 446 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 447 | lemma type_substitution_lemma: | 
| 32011 | 448 | "X \<noteq> Y \<Longrightarrow> X\<sharp>L \<Longrightarrow> M[X \<mapsto> N]\<^sub>\<tau>[Y \<mapsto> L]\<^sub>\<tau> = M[Y \<mapsto> L]\<^sub>\<tau>[X \<mapsto> N[Y \<mapsto> L]\<^sub>\<tau>]\<^sub>\<tau>" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 449 | by (nominal_induct M avoiding: X Y N L rule: ty.strong_induct) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 450 | (auto simp add: type_subst_fresh type_subst_identity) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 451 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 452 | lemma type_subst_rename: | 
| 32011 | 453 | "Y\<sharp>T \<Longrightarrow> ([(Y,X)]\<bullet>T)[Y \<mapsto> U]\<^sub>\<tau> = T[X \<mapsto> U]\<^sub>\<tau>" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 454 | by (nominal_induct T avoiding: X Y U rule: ty.strong_induct) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 455 | (simp_all add: fresh_atm calc_atm abs_fresh fresh_aux) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 456 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 457 | nominal_primrec | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 458 |   subst_tyb :: "binding \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> binding" ("_[_ \<mapsto> _]\<^sub>b" [100,100,100] 100)
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 459 | where | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 460 | "(TVarB X U)[Y \<mapsto> T]\<^sub>b = TVarB X (U[Y \<mapsto> T]\<^sub>\<tau>)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 461 | | "(VarB X U)[Y \<mapsto> T]\<^sub>b = VarB X (U[Y \<mapsto> T]\<^sub>\<tau>)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 462 | by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 463 | |
| 30986 
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
 Christian Urban <urbanc@in.tum.de> parents: 
30091diff
changeset | 464 | lemma binding_subst_fresh: | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 465 | fixes X::"tyvrs" | 
| 32011 | 466 | assumes "X\<sharp>a" | 
| 467 | and "X\<sharp>P" | |
| 468 | shows "X\<sharp>a[Y \<mapsto> P]\<^sub>b" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 469 | using assms | 
| 30986 
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
 Christian Urban <urbanc@in.tum.de> parents: 
30091diff
changeset | 470 | by (nominal_induct a rule: binding.strong_induct) | 
| 
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
 Christian Urban <urbanc@in.tum.de> parents: 
30091diff
changeset | 471 | (auto simp add: type_subst_fresh) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 472 | |
| 30986 
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
 Christian Urban <urbanc@in.tum.de> parents: 
30091diff
changeset | 473 | lemma binding_subst_identity: | 
| 32011 | 474 | shows "X\<sharp>B \<Longrightarrow> B[X \<mapsto> U]\<^sub>b = B" | 
| 30986 
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
 Christian Urban <urbanc@in.tum.de> parents: 
30091diff
changeset | 475 | by (induct B rule: binding.induct) | 
| 
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
 Christian Urban <urbanc@in.tum.de> parents: 
30091diff
changeset | 476 | (simp_all add: fresh_atm type_subst_identity) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 477 | |
| 22537 | 478 | consts | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 479 |   subst_tyc :: "env \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> env" ("_[_ \<mapsto> _]\<^sub>e" [100,100,100] 100)
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 480 | |
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 481 | primrec | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 482 | "([])[Y \<mapsto> T]\<^sub>e= []" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 483 | "(B#\<Gamma>)[Y \<mapsto> T]\<^sub>e = (B[Y \<mapsto> T]\<^sub>b)#(\<Gamma>[Y \<mapsto> T]\<^sub>e)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 484 | |
| 30986 
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
 Christian Urban <urbanc@in.tum.de> parents: 
30091diff
changeset | 485 | lemma ctxt_subst_fresh': | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 486 | fixes X::"tyvrs" | 
| 32011 | 487 | assumes "X\<sharp>\<Gamma>" | 
| 488 | and "X\<sharp>P" | |
| 489 | shows "X\<sharp>\<Gamma>[Y \<mapsto> P]\<^sub>e" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 490 | using assms | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 491 | by (induct \<Gamma>) | 
| 30986 
047fa04a9fe8
deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
 Christian Urban <urbanc@in.tum.de> parents: 
30091diff
changeset | 492 | (auto simp add: fresh_list_cons binding_subst_fresh) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 493 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 494 | lemma ctxt_subst_mem_TVarB: "TVarB X T \<in> set \<Gamma> \<Longrightarrow> TVarB X (T[Y \<mapsto> U]\<^sub>\<tau>) \<in> set (\<Gamma>[Y \<mapsto> U]\<^sub>e)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 495 | by (induct \<Gamma>) auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 496 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 497 | lemma ctxt_subst_mem_VarB: "VarB x T \<in> set \<Gamma> \<Longrightarrow> VarB x (T[Y \<mapsto> U]\<^sub>\<tau>) \<in> set (\<Gamma>[Y \<mapsto> U]\<^sub>e)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 498 | by (induct \<Gamma>) auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 499 | |
| 32011 | 500 | lemma ctxt_subst_identity: "X\<sharp>\<Gamma> \<Longrightarrow> \<Gamma>[X \<mapsto> U]\<^sub>e = \<Gamma>" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 501 | by (induct \<Gamma>) (simp_all add: fresh_list_cons binding_subst_identity) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 502 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 503 | lemma ctxt_subst_append: "(\<Delta> @ \<Gamma>)[X \<mapsto> T]\<^sub>e = \<Delta>[X \<mapsto> T]\<^sub>e @ \<Gamma>[X \<mapsto> T]\<^sub>e" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 504 | by (induct \<Delta>) simp_all | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 505 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 506 | nominal_primrec | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 507 |    subst_trm :: "trm \<Rightarrow> vrs \<Rightarrow> trm \<Rightarrow> trm"  ("_[_ \<mapsto> _]" [300, 0, 0] 300)
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 508 | where | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 509 | "(Var x)[y \<mapsto> t'] = (if x=y then t' else (Var x))" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 510 | | "(t1 \<cdot> t2)[y \<mapsto> t'] = t1[y \<mapsto> t'] \<cdot> t2[y \<mapsto> t']" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 511 | | "(t \<cdot>\<^sub>\<tau> T)[y \<mapsto> t'] = t[y \<mapsto> t'] \<cdot>\<^sub>\<tau> T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 512 | | "X\<sharp>(T,t') \<Longrightarrow> (\<lambda>X<:T. t)[y \<mapsto> t'] = (\<lambda>X<:T. t[y \<mapsto> t'])" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 513 | | "x\<sharp>(y,t') \<Longrightarrow> (\<lambda>x:T. t)[y \<mapsto> t'] = (\<lambda>x:T. t[y \<mapsto> t'])" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 514 | apply(finite_guess)+ | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 515 | apply(rule TrueI)+ | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 516 | apply(simp add: abs_fresh)+ | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 517 | apply(fresh_guess add: ty_vrs_fresh abs_fresh)+ | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 518 | done | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 519 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 520 | lemma subst_trm_fresh_tyvar: | 
| 32011 | 521 | fixes X::"tyvrs" | 
| 522 | shows "X\<sharp>t \<Longrightarrow> X\<sharp>u \<Longrightarrow> X\<sharp>t[x \<mapsto> u]" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 523 | by (nominal_induct t avoiding: x u rule: trm.strong_induct) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 524 | (auto simp add: trm.fresh abs_fresh) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 525 | |
| 32011 | 526 | lemma subst_trm_fresh_var: | 
| 527 | "x\<sharp>u \<Longrightarrow> x\<sharp>t[x \<mapsto> u]" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 528 | by (nominal_induct t avoiding: x u rule: trm.strong_induct) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 529 | (simp_all add: abs_fresh fresh_atm ty_vrs_fresh) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 530 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 531 | lemma subst_trm_eqvt[eqvt]: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 532 | fixes pi::"tyvrs prm" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 533 | and t::"trm" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 534 | shows "pi\<bullet>(t[x \<mapsto> u]) = (pi\<bullet>t)[(pi\<bullet>x) \<mapsto> (pi\<bullet>u)]" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 535 | by (nominal_induct t avoiding: x u rule: trm.strong_induct) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 536 | (perm_simp add: fresh_left)+ | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 537 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 538 | lemma subst_trm_eqvt'[eqvt]: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 539 | fixes pi::"vrs prm" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 540 | and t::"trm" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 541 | shows "pi\<bullet>(t[x \<mapsto> u]) = (pi\<bullet>t)[(pi\<bullet>x) \<mapsto> (pi\<bullet>u)]" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 542 | by (nominal_induct t avoiding: x u rule: trm.strong_induct) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 543 | (perm_simp add: fresh_left)+ | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 544 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 545 | lemma subst_trm_rename: | 
| 32011 | 546 | "y\<sharp>t \<Longrightarrow> ([(y, x)] \<bullet> t)[y \<mapsto> u] = t[x \<mapsto> u]" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 547 | by (nominal_induct t avoiding: x y u rule: trm.strong_induct) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 548 | (simp_all add: fresh_atm calc_atm abs_fresh fresh_aux ty_vrs_fresh perm_fresh_fresh) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 549 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 550 | nominal_primrec (freshness_context: "T2::ty") | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 551 |   subst_trm_ty :: "trm \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> trm"  ("_[_ \<mapsto>\<^sub>\<tau> _]" [300, 0, 0] 300)
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 552 | where | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 553 | "(Var x)[Y \<mapsto>\<^sub>\<tau> T2] = Var x" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 554 | | "(t1 \<cdot> t2)[Y \<mapsto>\<^sub>\<tau> T2] = t1[Y \<mapsto>\<^sub>\<tau> T2] \<cdot> t2[Y \<mapsto>\<^sub>\<tau> T2]" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 555 | | "(t1 \<cdot>\<^sub>\<tau> T)[Y \<mapsto>\<^sub>\<tau> T2] = t1[Y \<mapsto>\<^sub>\<tau> T2] \<cdot>\<^sub>\<tau> T[Y \<mapsto> T2]\<^sub>\<tau>" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 556 | | "X\<sharp>(Y,T,T2) \<Longrightarrow> (\<lambda>X<:T. t)[Y \<mapsto>\<^sub>\<tau> T2] = (\<lambda>X<:T[Y \<mapsto> T2]\<^sub>\<tau>. t[Y \<mapsto>\<^sub>\<tau> T2])" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 557 | | "(\<lambda>x:T. t)[Y \<mapsto>\<^sub>\<tau> T2] = (\<lambda>x:T[Y \<mapsto> T2]\<^sub>\<tau>. t[Y \<mapsto>\<^sub>\<tau> T2])" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 558 | apply(finite_guess)+ | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 559 | apply(rule TrueI)+ | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 560 | apply(simp add: abs_fresh ty_vrs_fresh)+ | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 561 | apply(simp add: type_subst_fresh) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 562 | apply(fresh_guess add: ty_vrs_fresh abs_fresh)+ | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 563 | done | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 564 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 565 | lemma subst_trm_ty_fresh: | 
| 32011 | 566 | fixes X::"tyvrs" | 
| 567 | shows "X\<sharp>t \<Longrightarrow> X\<sharp>T \<Longrightarrow> X\<sharp>t[Y \<mapsto>\<^sub>\<tau> T]" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 568 | by (nominal_induct t avoiding: Y T rule: trm.strong_induct) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 569 | (auto simp add: abs_fresh type_subst_fresh) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 570 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 571 | lemma subst_trm_ty_fresh': | 
| 32011 | 572 | "X\<sharp>T \<Longrightarrow> X\<sharp>t[X \<mapsto>\<^sub>\<tau> T]" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 573 | by (nominal_induct t avoiding: X T rule: trm.strong_induct) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 574 | (simp_all add: abs_fresh fresh_type_subst_fresh fresh_atm) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 575 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 576 | lemma subst_trm_ty_eqvt[eqvt]: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 577 | fixes pi::"tyvrs prm" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 578 | and t::"trm" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 579 | shows "pi\<bullet>(t[X \<mapsto>\<^sub>\<tau> T]) = (pi\<bullet>t)[(pi\<bullet>X) \<mapsto>\<^sub>\<tau> (pi\<bullet>T)]" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 580 | by (nominal_induct t avoiding: X T rule: trm.strong_induct) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 581 | (perm_simp add: fresh_bij subst_eqvt)+ | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 582 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 583 | lemma subst_trm_ty_eqvt'[eqvt]: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 584 | fixes pi::"vrs prm" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 585 | and t::"trm" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 586 | shows "pi\<bullet>(t[X \<mapsto>\<^sub>\<tau> T]) = (pi\<bullet>t)[(pi\<bullet>X) \<mapsto>\<^sub>\<tau> (pi\<bullet>T)]" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 587 | by (nominal_induct t avoiding: X T rule: trm.strong_induct) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 588 | (perm_simp add: fresh_left subst_eqvt')+ | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 589 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 590 | lemma subst_trm_ty_rename: | 
| 32011 | 591 | "Y\<sharp>t \<Longrightarrow> ([(Y, X)] \<bullet> t)[Y \<mapsto>\<^sub>\<tau> U] = t[X \<mapsto>\<^sub>\<tau> U]" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 592 | by (nominal_induct t avoiding: X Y U rule: trm.strong_induct) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 593 | (simp_all add: fresh_atm calc_atm abs_fresh fresh_aux type_subst_rename) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 594 | |
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 595 | section {* Subtyping-Relation *}
 | 
| 18246 | 596 | |
| 18650 | 597 | text {* The definition for the subtyping-relation follows quite closely what is written 
 | 
| 598 | in the POPLmark-paper, except for the premises dealing with well-formed contexts and | |
| 599 |   the freshness constraint @{term "X\<sharp>\<Gamma>"} in the @{text "S_Forall"}-rule. (The freshness
 | |
| 600 |   constraint is specific to the \emph{nominal approach}. Note, however, that the constraint
 | |
| 601 |   does \emph{not} make the subtyping-relation ``partial"\ldots because we work over
 | |
| 602 | $\alpha$-equivalence classes.) *} | |
| 18628 | 603 | |
| 23760 | 604 | inductive | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 605 |   subtype_of :: "env \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool"   ("_\<turnstile>_<:_" [100,100,100] 100)
 | 
| 22436 | 606 | where | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 607 | SA_Top[intro]: "\<lbrakk>\<turnstile> \<Gamma> ok; S closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> S <: Top" | 
| 32011 | 608 | | SA_refl_TVar[intro]: "\<lbrakk>\<turnstile> \<Gamma> ok; X \<in> ty_dom \<Gamma>\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> Tvar X <: Tvar X" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 609 | | SA_trans_TVar[intro]: "\<lbrakk>(TVarB X S) \<in> set \<Gamma>; \<Gamma> \<turnstile> S <: T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (Tvar X) <: T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 610 | | SA_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)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 611 | | SA_all[intro]: "\<lbrakk>\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1; ((TVarB 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)" | 
| 22537 | 612 | |
| 613 | lemma subtype_implies_ok: | |
| 614 | fixes X::"tyvrs" | |
| 615 | assumes a: "\<Gamma> \<turnstile> S <: T" | |
| 616 | shows "\<turnstile> \<Gamma> ok" | |
| 617 | using a by (induct) (auto) | |
| 18246 | 618 | |
| 619 | lemma subtype_implies_closed: | |
| 620 | assumes a: "\<Gamma> \<turnstile> S <: T" | |
| 621 | shows "S closed_in \<Gamma> \<and> T closed_in \<Gamma>" | |
| 622 | using a | |
| 623 | proof (induct) | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 624 | case (SA_Top \<Gamma> S) | 
| 18424 | 625 | have "Top closed_in \<Gamma>" by (simp add: closed_in_def ty.supp) | 
| 18246 | 626 | moreover | 
| 627 | have "S closed_in \<Gamma>" by fact | |
| 628 | ultimately show "S closed_in \<Gamma> \<and> Top closed_in \<Gamma>" by simp | |
| 629 | next | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 630 | case (SA_trans_TVar X S \<Gamma> T) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 631 | have "(TVarB X S)\<in>set \<Gamma>" by fact | 
| 32011 | 632 | hence "X \<in> ty_dom \<Gamma>" by (rule ty_dom_inclusion) | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 633 | hence "(Tvar X) closed_in \<Gamma>" by (simp add: closed_in_def ty.supp supp_atm) | 
| 18246 | 634 | moreover | 
| 635 | have "S closed_in \<Gamma> \<and> T closed_in \<Gamma>" by fact | |
| 636 | hence "T closed_in \<Gamma>" by force | |
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 637 | ultimately show "(Tvar X) closed_in \<Gamma> \<and> T closed_in \<Gamma>" by simp | 
| 18424 | 638 | qed (auto simp add: closed_in_def ty.supp supp_atm abs_supp) | 
| 18246 | 639 | |
| 640 | lemma subtype_implies_fresh: | |
| 641 | fixes X::"tyvrs" | |
| 642 | assumes a1: "\<Gamma> \<turnstile> S <: T" | |
| 643 | and a2: "X\<sharp>\<Gamma>" | |
| 18424 | 644 | shows "X\<sharp>S \<and> X\<sharp>T" | 
| 18246 | 645 | proof - | 
| 646 | from a1 have "\<turnstile> \<Gamma> ok" by (rule subtype_implies_ok) | |
| 32011 | 647 | with a2 have "X\<sharp>ty_dom(\<Gamma>)" by (simp add: fresh_dom) | 
| 18424 | 648 | moreover | 
| 18246 | 649 | from a1 have "S closed_in \<Gamma> \<and> T closed_in \<Gamma>" by (rule subtype_implies_closed) | 
| 32011 | 650 | hence "supp S \<subseteq> ((supp (ty_dom \<Gamma>))::tyvrs set)" | 
| 651 | and "supp T \<subseteq> ((supp (ty_dom \<Gamma>))::tyvrs set)" by (simp_all add: ty_dom_supp closed_in_def) | |
| 18424 | 652 | ultimately show "X\<sharp>S \<and> X\<sharp>T" by (force simp add: supp_prod fresh_def) | 
| 18246 | 653 | qed | 
| 654 | ||
| 32011 | 655 | lemma valid_ty_dom_fresh: | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 656 | fixes X::"tyvrs" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 657 | assumes valid: "\<turnstile> \<Gamma> ok" | 
| 32011 | 658 | shows "X\<sharp>(ty_dom \<Gamma>) = X\<sharp>\<Gamma>" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 659 | using valid | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 660 | apply induct | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 661 | apply (simp add: fresh_list_nil fresh_set_empty) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 662 | apply (simp_all add: binding.fresh fresh_list_cons | 
| 32011 | 663 | fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_doms fresh_atm) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 664 | apply (auto simp add: closed_in_fresh) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 665 | done | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 666 | |
| 22730 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22542diff
changeset | 667 | equivariance subtype_of | 
| 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22542diff
changeset | 668 | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 669 | nominal_inductive subtype_of | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 670 | apply (simp_all add: abs_fresh) | 
| 32011 | 671 | apply (fastsimp simp add: valid_ty_dom_fresh dest: subtype_implies_ok) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 672 | apply (force simp add: closed_in_fresh dest: subtype_implies_closed subtype_implies_ok)+ | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 673 | done | 
| 18246 | 674 | |
| 18621 | 675 | section {* Reflexivity of Subtyping *}
 | 
| 18246 | 676 | |
| 677 | lemma subtype_reflexivity: | |
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 678 | assumes a: "\<turnstile> \<Gamma> ok" | 
| 18424 | 679 | and b: "T closed_in \<Gamma>" | 
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 680 | shows "\<Gamma> \<turnstile> T <: T" | 
| 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 681 | using a b | 
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26091diff
changeset | 682 | 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 | 683 | case (Forall X T\<^isub>1 T\<^isub>2) | 
| 18747 | 684 | 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 | 
| 685 | 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 | 686 | have fresh_cond: "X\<sharp>\<Gamma>" by fact | 
| 32011 | 687 | hence fresh_ty_dom: "X\<sharp>(ty_dom \<Gamma>)" by (simp add: fresh_dom) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 688 | have "(\<forall>X<:T\<^isub>2. T\<^isub>1) closed_in \<Gamma>" by fact | 
| 35438 | 689 | hence closed\<^isub>T\<^isub>2: "T\<^isub>2 closed_in \<Gamma>" and closed\<^isub>T\<^isub>1: "T\<^isub>1 closed_in ((TVarB X T\<^isub>2)#\<Gamma>)" | 
| 18424 | 690 | by (auto simp add: closed_in_def ty.supp abs_supp) | 
| 691 | have ok: "\<turnstile> \<Gamma> ok" by fact | |
| 35438 | 692 | hence ok': "\<turnstile> ((TVarB X T\<^isub>2)#\<Gamma>) ok" using closed\<^isub>T\<^isub>2 fresh_ty_dom by simp | 
| 693 | 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 | 694 | moreover | 
| 35438 | 695 | have "((TVarB 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 | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 696 | ultimately show "\<Gamma> \<turnstile> (\<forall>X<:T\<^isub>2. T\<^isub>1) <: (\<forall>X<:T\<^isub>2. T\<^isub>1)" using fresh_cond | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 697 | by (simp add: subtype_of.SA_all) | 
| 18246 | 698 | qed (auto simp add: closed_in_def ty.supp supp_atm) | 
| 699 | ||
| 18621 | 700 | lemma subtype_reflexivity_semiautomated: | 
| 18305 
a780f9c1538b
changed everything until the interesting transitivity_narrowing
 urbanc parents: 
18269diff
changeset | 701 | assumes a: "\<turnstile> \<Gamma> ok" | 
| 
a780f9c1538b
changed everything until the interesting transitivity_narrowing
 urbanc parents: 
18269diff
changeset | 702 | and b: "T closed_in \<Gamma>" | 
| 
a780f9c1538b
changed everything until the interesting transitivity_narrowing
 urbanc parents: 
18269diff
changeset | 703 | shows "\<Gamma> \<turnstile> T <: T" | 
| 
a780f9c1538b
changed everything until the interesting transitivity_narrowing
 urbanc parents: 
18269diff
changeset | 704 | using a b | 
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
26091diff
changeset | 705 | apply(nominal_induct T avoiding: \<Gamma> rule: ty.strong_induct) | 
| 18747 | 706 | 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 | 707 |   --{* Too bad that this instantiation cannot be found automatically by
 | 
| 18621 | 708 |   \isakeyword{auto}; \isakeyword{blast} would find it if we had not used 
 | 
| 18628 | 709 |   an explicit definition for @{text "closed_in_def"}. *}
 | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 710 | apply(drule_tac x="(TVarB tyvrs ty2)#\<Gamma>" in meta_spec) | 
| 32011 | 711 | apply(force dest: fresh_dom simp add: closed_in_def) | 
| 18246 | 712 | done | 
| 713 | ||
| 18628 | 714 | section {* Weakening *}
 | 
| 18246 | 715 | |
| 18628 | 716 | text {* In order to prove weakening we introduce the notion of a type-context extending 
 | 
| 717 | another. This generalization seems to make the proof for weakening to be | |
| 718 | smoother than if we had strictly adhered to the version in the POPLmark-paper. *} | |
| 18246 | 719 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
34915diff
changeset | 720 | definition extends :: "env \<Rightarrow> env \<Rightarrow> bool" ("_ extends _" [100,100] 100) where
 | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 721 | "\<Delta> extends \<Gamma> \<equiv> \<forall>X Q. (TVarB X Q)\<in>set \<Gamma> \<longrightarrow> (TVarB X Q)\<in>set \<Delta>" | 
| 18246 | 722 | |
| 32011 | 723 | lemma extends_ty_dom: | 
| 18246 | 724 | assumes a: "\<Delta> extends \<Gamma>" | 
| 32011 | 725 | shows "ty_dom \<Gamma> \<subseteq> ty_dom \<Delta>" | 
| 18246 | 726 | using a | 
| 727 | apply (auto simp add: extends_def) | |
| 32011 | 728 | apply (drule ty_dom_existence) | 
| 729 | apply (force simp add: ty_dom_inclusion) | |
| 18246 | 730 | done | 
| 731 | ||
| 732 | lemma extends_closed: | |
| 733 | assumes a1: "T closed_in \<Gamma>" | |
| 734 | and a2: "\<Delta> extends \<Gamma>" | |
| 735 | shows "T closed_in \<Delta>" | |
| 736 | using a1 a2 | |
| 32011 | 737 | by (auto dest: extends_ty_dom simp add: closed_in_def) | 
| 18246 | 738 | |
| 18424 | 739 | lemma extends_memb: | 
| 740 | assumes a: "\<Delta> extends \<Gamma>" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 741 | and b: "(TVarB X T) \<in> set \<Gamma>" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 742 | shows "(TVarB X T) \<in> set \<Delta>" | 
| 18424 | 743 | using a b by (simp add: extends_def) | 
| 744 | ||
| 18246 | 745 | lemma weakening: | 
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 746 | assumes a: "\<Gamma> \<turnstile> S <: T" | 
| 18424 | 747 | and b: "\<turnstile> \<Delta> ok" | 
| 748 | and c: "\<Delta> extends \<Gamma>" | |
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 749 | shows "\<Delta> \<turnstile> S <: T" | 
| 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 750 | using a b c | 
| 22537 | 751 | proof (nominal_induct \<Gamma> S T avoiding: \<Delta> rule: subtype_of.strong_induct) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 752 | case (SA_Top \<Gamma> S) | 
| 18246 | 753 | 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 | 754 | have "\<turnstile> \<Delta> ok" by fact | 
| 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 755 | moreover | 
| 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 756 | have "\<Delta> extends \<Gamma>" by fact | 
| 18424 | 757 | 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 | 758 | ultimately show "\<Delta> \<turnstile> S <: Top" by force | 
| 18246 | 759 | next | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 760 | case (SA_trans_TVar X S \<Gamma> T) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 761 | have lh_drv_prem: "(TVarB X S) \<in> set \<Gamma>" by fact | 
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 762 | 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 | 763 | have ok: "\<turnstile> \<Delta> ok" by fact | 
| 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 764 | have extends: "\<Delta> extends \<Gamma>" by fact | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 765 | have "(TVarB 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 | 766 | moreover | 
| 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 767 | 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 | 768 | ultimately show "\<Delta> \<turnstile> Tvar X <: T" using ok by force | 
| 18246 | 769 | next | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 770 | case (SA_refl_TVar \<Gamma> X) | 
| 32011 | 771 | have lh_drv_prem: "X \<in> ty_dom \<Gamma>" by fact | 
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 772 | have "\<turnstile> \<Delta> ok" by fact | 
| 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 773 | moreover | 
| 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 774 | have "\<Delta> extends \<Gamma>" by fact | 
| 32011 | 775 | hence "X \<in> ty_dom \<Delta>" using lh_drv_prem by (force dest: extends_ty_dom) | 
| 18577 
a636846a02c7
added more documentation; will now try out a modification
 urbanc parents: 
18424diff
changeset | 776 | ultimately show "\<Delta> \<turnstile> Tvar X <: Tvar X" by force | 
| 18246 | 777 | next | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 778 | case (SA_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 | 779 | next | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 780 | case (SA_all \<Gamma> T\<^isub>1 S\<^isub>1 X S\<^isub>2 T\<^isub>2) | 
| 18424 | 781 | have fresh_cond: "X\<sharp>\<Delta>" by fact | 
| 32011 | 782 | hence fresh_dom: "X\<sharp>(ty_dom \<Delta>)" by (simp add: fresh_dom) | 
| 18424 | 783 | have ih\<^isub>1: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends \<Gamma> \<Longrightarrow> \<Delta> \<turnstile> T\<^isub>1 <: S\<^isub>1" by fact | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 784 | have ih\<^isub>2: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends ((TVarB X T\<^isub>1)#\<Gamma>) \<Longrightarrow> \<Delta> \<turnstile> S\<^isub>2 <: T\<^isub>2" by fact | 
| 18424 | 785 | have lh_drv_prem: "\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1" by fact | 
| 35438 | 786 | 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 | 787 | have ok: "\<turnstile> \<Delta> ok" by fact | 
| 18424 | 788 | have ext: "\<Delta> extends \<Gamma>" by fact | 
| 35438 | 789 | have "T\<^isub>1 closed_in \<Delta>" using ext closed\<^isub>T\<^isub>1 by (simp only: extends_closed) | 
| 32011 | 790 | hence "\<turnstile> ((TVarB X T\<^isub>1)#\<Delta>) ok" using fresh_dom ok by force | 
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 791 | moreover | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 792 | have "((TVarB X T\<^isub>1)#\<Delta>) extends ((TVarB X T\<^isub>1)#\<Gamma>)" using ext by (force simp add: extends_def) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 793 | ultimately have "((TVarB 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 | 794 | moreover | 
| 18424 | 795 | have "\<Delta> \<turnstile> T\<^isub>1 <: S\<^isub>1" using ok ext ih\<^isub>1 by simp | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 796 | 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: SA_all) | 
| 18246 | 797 | qed | 
| 798 | ||
| 18650 | 799 | text {* In fact all ``non-binding" cases can be solved automatically: *}
 | 
| 18246 | 800 | |
| 18628 | 801 | lemma weakening_more_automated: | 
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 802 | assumes a: "\<Gamma> \<turnstile> S <: T" | 
| 18424 | 803 | and b: "\<turnstile> \<Delta> ok" | 
| 804 | and c: "\<Delta> extends \<Gamma>" | |
| 18353 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 805 | shows "\<Delta> \<turnstile> S <: T" | 
| 
4dd468ccfdf7
transitivity should be now in a reasonable state. But
 urbanc parents: 
18306diff
changeset | 806 | using a b c | 
| 22537 | 807 | proof (nominal_induct \<Gamma> S T avoiding: \<Delta> rule: subtype_of.strong_induct) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 808 | case (SA_all \<Gamma> T\<^isub>1 S\<^isub>1 X S\<^isub>2 T\<^isub>2) | 
| 18424 | 809 | have fresh_cond: "X\<sharp>\<Delta>" by fact | 
| 32011 | 810 | hence fresh_dom: "X\<sharp>(ty_dom \<Delta>)" by (simp add: fresh_dom) | 
| 18424 | 811 | have ih\<^isub>1: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends \<Gamma> \<Longrightarrow> \<Delta> \<turnstile> T\<^isub>1 <: S\<^isub>1" by fact | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 812 | have ih\<^isub>2: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends ((TVarB X T\<^isub>1)#\<Gamma>) \<Longrightarrow> \<Delta> \<turnstile> S\<^isub>2 <: T\<^isub>2" by fact | 
| 18424 | 813 | have lh_drv_prem: "\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1" by fact | 
| 35438 | 814 | 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 | 815 | have ok: "\<turnstile> \<Delta> ok" by fact | 
| 18424 | 816 | have ext: "\<Delta> extends \<Gamma>" by fact | 
| 35438 | 817 | have "T\<^isub>1 closed_in \<Delta>" using ext closed\<^isub>T\<^isub>1 by (simp only: extends_closed) | 
| 32011 | 818 | hence "\<turnstile> ((TVarB X T\<^isub>1)#\<Delta>) ok" using fresh_dom ok by force | 
| 18628 | 819 | moreover | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 820 | have "((TVarB X T\<^isub>1)#\<Delta>) extends ((TVarB X T\<^isub>1)#\<Gamma>)" using ext by (force simp add: extends_def) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 821 | ultimately have "((TVarB 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 | 822 | moreover | 
| 18424 | 823 | have "\<Delta> \<turnstile> T\<^isub>1 <: S\<^isub>1" using ok ext ih\<^isub>1 by simp | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 824 | 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: SA_all) | 
| 32011 | 825 | qed (blast intro: extends_closed extends_memb dest: extends_ty_dom)+ | 
| 18246 | 826 | |
| 18628 | 827 | section {* Transitivity and Narrowing *}
 | 
| 828 | ||
| 18650 | 829 | text {* Some inversion lemmas that are needed in the transitivity and narrowing proof.*}
 | 
| 830 | ||
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 831 | declare ty.inject [simp add] | 
| 18650 | 832 | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 833 | inductive_cases S_TopE: "\<Gamma> \<turnstile> Top <: T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 834 | inductive_cases S_ArrowE_left: "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 835 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 836 | declare ty.inject [simp del] | 
| 18650 | 837 | |
| 838 | lemma S_ForallE_left: | |
| 32011 | 839 | shows "\<lbrakk>\<Gamma> \<turnstile> (\<forall>X<:S\<^isub>1. S\<^isub>2) <: T; X\<sharp>\<Gamma>; X\<sharp>S\<^isub>1; X\<sharp>T\<rbrakk> | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 840 | \<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> ((TVarB X T\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2)" | 
| 32011 | 841 | apply(erule subtype_of.strong_cases[where X="X"]) | 
| 842 | apply(auto simp add: abs_fresh ty.inject alpha) | |
| 843 | done | |
| 18650 | 844 | |
| 845 | text {* Next we prove the transitivity and narrowing for the subtyping-relation. 
 | |
| 18621 | 846 | The POPLmark-paper says the following: | 
| 847 | ||
| 18650 | 848 | \begin{quote}
 | 
| 18621 | 849 | \begin{lemma}[Transitivity and Narrowing] \
 | 
| 850 | \begin{enumerate}
 | |
| 851 | \item If @{term "\<Gamma> \<turnstile> S<:Q"} and @{term "\<Gamma> \<turnstile> Q<:T"}, then @{term "\<Gamma> \<turnstile> S<:T"}.
 | |
| 852 | \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"}.
 | |
| 853 | \end{enumerate}
 | |
| 854 | \end{lemma}
 | |
| 855 | ||
| 856 | The two parts are proved simultaneously, by induction on the size | |
| 857 | of @{term "Q"}.  The argument for part (2) assumes that part (1) has 
 | |
| 858 | been established already for the @{term "Q"} in question; part (1) uses 
 | |
| 859 | part (2) only for strictly smaller @{term "Q"}.
 | |
| 18650 | 860 | \end{quote}
 | 
| 18621 | 861 | |
| 862 | For the induction on the size of @{term "Q"}, we use the induction-rule 
 | |
| 863 | @{text "measure_induct_rule"}:
 | |
| 864 | ||
| 865 | \begin{center}
 | |
| 866 | @{thm measure_induct_rule[of "size_ty",no_vars]}
 | |
| 867 | \end{center}
 | |
| 18410 | 868 | |
| 18628 | 869 | That means in order to show a property @{term "P a"} for all @{term "a"}, 
 | 
| 18650 | 870 | the induct-rule requires to prove that for all @{term x} @{term "P x"} holds using the 
 | 
| 18621 | 871 | assumption that for all @{term y} whose size is strictly smaller than 
 | 
| 872 | 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 | 873 | |
| 18621 | 874 | lemma | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 875 | shows subtype_transitivity: "\<Gamma>\<turnstile>S<:Q \<Longrightarrow> \<Gamma>\<turnstile>Q<:T \<Longrightarrow> \<Gamma>\<turnstile>S<:T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 876 | and subtype_narrow: "(\<Delta>@[(TVarB X Q)]@\<Gamma>)\<turnstile>M<:N \<Longrightarrow> \<Gamma>\<turnstile>P<:Q \<Longrightarrow> (\<Delta>@[(TVarB X P)]@\<Gamma>)\<turnstile>M<:N" | 
| 20503 | 877 | proof (induct Q arbitrary: \<Gamma> S T \<Delta> X P M N taking: "size_ty" rule: measure_induct_rule) | 
| 18621 | 878 | case (less Q) | 
| 879 | have IH_trans: | |
| 880 | "\<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 | |
| 881 | have IH_narrow: | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 882 | "\<And>Q' \<Delta> \<Gamma> X M N P. \<lbrakk>size_ty Q' < size_ty Q; (\<Delta>@[(TVarB X Q')]@\<Gamma>)\<turnstile>M<:N; \<Gamma>\<turnstile>P<:Q'\<rbrakk> | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 883 | \<Longrightarrow> (\<Delta>@[(TVarB X P)]@\<Gamma>)\<turnstile>M<:N" by fact | 
| 32011 | 884 | |
| 885 |   { fix \<Gamma> S T
 | |
| 886 | have "\<lbrakk>\<Gamma> \<turnstile> S <: Q; \<Gamma> \<turnstile> Q <: T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> S <: T" | |
| 887 | proof (induct \<Gamma> S Q\<equiv>Q rule: subtype_of.induct) | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 888 | case (SA_Top \<Gamma> S) | 
| 32011 | 889 | then have rh_drv: "\<Gamma> \<turnstile> Top <: T" by simp | 
| 890 | then have T_inst: "T = Top" by (auto elim: S_TopE) | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 891 | from `\<turnstile> \<Gamma> ok` and `S closed_in \<Gamma>` | 
| 32011 | 892 | have "\<Gamma> \<turnstile> S <: Top" by auto | 
| 893 | then show "\<Gamma> \<turnstile> S <: T" using T_inst by simp | |
| 18246 | 894 | next | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 895 | case (SA_trans_TVar Y U \<Gamma>) | 
| 32011 | 896 | then have IH_inner: "\<Gamma> \<turnstile> U <: T" by simp | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 897 | have "(TVarB Y U) \<in> set \<Gamma>" by fact | 
| 32011 | 898 | with IH_inner show "\<Gamma> \<turnstile> Tvar Y <: T" by auto | 
| 18246 | 899 | next | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 900 | case (SA_refl_TVar \<Gamma> X) | 
| 32011 | 901 | then show "\<Gamma> \<turnstile> Tvar X <: T" by simp | 
| 18246 | 902 | next | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 903 | case (SA_arrow \<Gamma> Q\<^isub>1 S\<^isub>1 S\<^isub>2 Q\<^isub>2) | 
| 32011 | 904 | then have rh_drv: "\<Gamma> \<turnstile> Q\<^isub>1 \<rightarrow> Q\<^isub>2 <: T" by simp | 
| 18621 | 905 | from `Q\<^isub>1 \<rightarrow> Q\<^isub>2 = Q` | 
| 35438 | 906 | 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 | 
| 18621 | 907 | have lh_drv_prm\<^isub>1: "\<Gamma> \<turnstile> Q\<^isub>1 <: S\<^isub>1" by fact | 
| 908 | have lh_drv_prm\<^isub>2: "\<Gamma> \<turnstile> S\<^isub>2 <: Q\<^isub>2" by fact | |
| 909 | 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)" | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 910 | by (auto elim: S_ArrowE_left) | 
| 18621 | 911 | moreover | 
| 912 | have "S\<^isub>1 closed_in \<Gamma>" and "S\<^isub>2 closed_in \<Gamma>" | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 913 | using lh_drv_prm\<^isub>1 lh_drv_prm\<^isub>2 by (simp_all add: subtype_implies_closed) | 
| 18621 | 914 | 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 | 915 | moreover | 
| 18424 | 916 | 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 | 917 | moreover | 
| 32011 | 918 |       { 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"
 | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 919 | then obtain T\<^isub>1 T\<^isub>2 | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 920 | where T_inst: "T = T\<^isub>1 \<rightarrow> T\<^isub>2" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 921 | and rh_drv_prm\<^isub>1: "\<Gamma> \<turnstile> T\<^isub>1 <: Q\<^isub>1" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 922 | and rh_drv_prm\<^isub>2: "\<Gamma> \<turnstile> Q\<^isub>2 <: T\<^isub>2" by force | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 923 | from IH_trans[of "Q\<^isub>1"] | 
| 35438 | 924 | 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 | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 925 | moreover | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 926 | from IH_trans[of "Q\<^isub>2"] | 
| 35438 | 927 | 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 | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 928 | ultimately have "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T\<^isub>1 \<rightarrow> T\<^isub>2" by auto | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 929 | then have "\<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 | 930 | } | 
| 18621 | 931 | ultimately show "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T" by blast | 
| 18246 | 932 | next | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 933 | case (SA_all \<Gamma> Q\<^isub>1 S\<^isub>1 X S\<^isub>2 Q\<^isub>2) | 
| 32011 | 934 | then have rh_drv: "\<Gamma> \<turnstile> (\<forall>X<:Q\<^isub>1. Q\<^isub>2) <: T" by simp | 
| 18621 | 935 | have lh_drv_prm\<^isub>1: "\<Gamma> \<turnstile> Q\<^isub>1 <: S\<^isub>1" by fact | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 936 | have lh_drv_prm\<^isub>2: "((TVarB X Q\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: Q\<^isub>2" by fact | 
| 32011 | 937 | then have "X\<sharp>\<Gamma>" by (force dest: subtype_implies_ok simp add: valid_ty_dom_fresh) | 
| 938 | then have fresh_cond: "X\<sharp>\<Gamma>" "X\<sharp>Q\<^isub>1" "X\<sharp>T" using rh_drv lh_drv_prm\<^isub>1 | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 939 | by (simp_all add: subtype_implies_fresh) | 
| 18621 | 940 | from rh_drv | 
| 32011 | 941 | have "T = Top \<or> | 
| 942 | (\<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> ((TVarB X T\<^isub>1)#\<Gamma>) \<turnstile> Q\<^isub>2 <: T\<^isub>2)" | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 943 | using fresh_cond by (simp add: S_ForallE_left) | 
| 18621 | 944 | moreover | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 945 | have "S\<^isub>1 closed_in \<Gamma>" and "S\<^isub>2 closed_in ((TVarB X Q\<^isub>1)#\<Gamma>)" | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 946 | using lh_drv_prm\<^isub>1 lh_drv_prm\<^isub>2 by (simp_all add: subtype_implies_closed) | 
| 32011 | 947 | then have "(\<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 | 948 | moreover | 
| 18424 | 949 | 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 | 950 | moreover | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 951 |       { 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> ((TVarB X T\<^isub>1)#\<Gamma>)\<turnstile>Q\<^isub>2<:T\<^isub>2"
 | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 952 | then obtain T\<^isub>1 T\<^isub>2 | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 953 | where T_inst: "T = (\<forall>X<:T\<^isub>1. T\<^isub>2)" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 954 | and rh_drv_prm\<^isub>1: "\<Gamma> \<turnstile> T\<^isub>1 <: Q\<^isub>1" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 955 | and rh_drv_prm\<^isub>2:"((TVarB X T\<^isub>1)#\<Gamma>) \<turnstile> Q\<^isub>2 <: T\<^isub>2" by force | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 956 | have "(\<forall>X<:Q\<^isub>1. Q\<^isub>2) = Q" by fact | 
| 35438 | 957 | then have Q\<^isub>1\<^isub>2_less: "size_ty Q\<^isub>1 < size_ty Q" "size_ty Q\<^isub>2 < size_ty Q" | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 958 | using fresh_cond by auto | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 959 | from IH_trans[of "Q\<^isub>1"] | 
| 35438 | 960 | 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 | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 961 | moreover | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 962 | from IH_narrow[of "Q\<^isub>1" "[]"] | 
| 35438 | 963 | have "((TVarB 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 | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 964 | with IH_trans[of "Q\<^isub>2"] | 
| 35438 | 965 | have "((TVarB 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 | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 966 | ultimately have "\<Gamma> \<turnstile> (\<forall>X<:S\<^isub>1. S\<^isub>2) <: (\<forall>X<:T\<^isub>1. T\<^isub>2)" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 967 | using fresh_cond by (simp add: subtype_of.SA_all) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 968 | 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 | 969 | } | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 970 | ultimately show "\<Gamma> \<turnstile> (\<forall>X<:S\<^isub>1. S\<^isub>2) <: T" by blast | 
| 18246 | 971 | qed | 
| 32011 | 972 | } note transitivity_lemma = this | 
| 18246 | 973 | |
| 18621 | 974 |   { --{* The transitivity proof is now by the auxiliary lemma. *}
 | 
| 975 | case 1 | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 976 | from `\<Gamma> \<turnstile> S <: Q` and `\<Gamma> \<turnstile> Q <: T` | 
| 32011 | 977 | show "\<Gamma> \<turnstile> S <: T" by (rule transitivity_lemma) | 
| 18621 | 978 | next | 
| 979 | case 2 | |
| 32011 | 980 | from `(\<Delta>@[(TVarB X Q)]@\<Gamma>) \<turnstile> M <: N` | 
| 981 | and `\<Gamma> \<turnstile> P<:Q` | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 982 | show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> M <: N" | 
| 34915 | 983 | proof (induct "\<Delta>@[(TVarB X Q)]@\<Gamma>" M N arbitrary: \<Gamma> X \<Delta> rule: subtype_of.induct) | 
| 984 | case (SA_Top S \<Gamma> X \<Delta>) | |
| 985 | from `\<Gamma> \<turnstile> P <: Q` | |
| 986 | have "P closed_in \<Gamma>" by (simp add: subtype_implies_closed) | |
| 987 | with `\<turnstile> (\<Delta>@[(TVarB X Q)]@\<Gamma>) ok` have "\<turnstile> (\<Delta>@[(TVarB X P)]@\<Gamma>) ok" | |
| 988 | by (simp add: replace_type) | |
| 18412 | 989 | moreover | 
| 34915 | 990 | from `S closed_in (\<Delta>@[(TVarB X Q)]@\<Gamma>)` have "S closed_in (\<Delta>@[(TVarB X P)]@\<Gamma>)" | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 991 | by (simp add: closed_in_def doms_append) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 992 | ultimately show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> S <: Top" by (simp add: subtype_of.SA_Top) | 
| 18246 | 993 | next | 
| 34915 | 994 | case (SA_trans_TVar Y S N \<Gamma> X \<Delta>) | 
| 32011 | 995 | then have IH_inner: "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> S <: N" | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 996 | and lh_drv_prm: "(TVarB Y S) \<in> set (\<Delta>@[(TVarB X Q)]@\<Gamma>)" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 997 | and rh_drv: "\<Gamma> \<turnstile> P<:Q" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 998 | and ok\<^isub>Q: "\<turnstile> (\<Delta>@[(TVarB X Q)]@\<Gamma>) ok" by (simp_all add: subtype_implies_ok) | 
| 32011 | 999 | then have ok\<^isub>P: "\<turnstile> (\<Delta>@[(TVarB X P)]@\<Gamma>) ok" by (simp add: subtype_implies_ok) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1000 | show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Tvar Y <: N" | 
| 18621 | 1001 | proof (cases "X=Y") | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 1002 | case False | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 1003 | have "X\<noteq>Y" by fact | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 1004 | hence "(TVarB Y S)\<in>set (\<Delta>@[(TVarB X P)]@\<Gamma>)" using lh_drv_prm by (simp add:binding.inject) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 1005 | with IH_inner show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Tvar Y <: N" by (simp add: subtype_of.SA_trans_TVar) | 
| 18621 | 1006 | next | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 1007 | case True | 
| 35438 | 1008 | have memb\<^isub>X\<^isub>Q: "(TVarB X Q)\<in>set (\<Delta>@[(TVarB X Q)]@\<Gamma>)" by simp | 
| 1009 | have memb\<^isub>X\<^isub>P: "(TVarB X P)\<in>set (\<Delta>@[(TVarB X P)]@\<Gamma>)" by simp | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 1010 | have eq: "X=Y" by fact | 
| 35438 | 1011 | hence "S=Q" using ok\<^isub>Q lh_drv_prm memb\<^isub>X\<^isub>Q by (simp only: uniqueness_of_ctxt) | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 1012 | hence "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Q <: N" using IH_inner by simp | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 1013 | moreover | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 1014 | have "(\<Delta>@[(TVarB X P)]@\<Gamma>) extends \<Gamma>" by (simp add: extends_def) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 1015 | hence "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> P <: Q" using rh_drv ok\<^isub>P by (simp only: weakening) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 1016 | ultimately have "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> P <: N" by (simp add: transitivity_lemma) | 
| 35438 | 1017 | then show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Tvar Y <: N" using memb\<^isub>X\<^isub>P eq by auto | 
| 18621 | 1018 | qed | 
| 18246 | 1019 | next | 
| 34915 | 1020 | case (SA_refl_TVar Y \<Gamma> X \<Delta>) | 
| 1021 | from `\<Gamma> \<turnstile> P <: Q` | |
| 1022 | have "P closed_in \<Gamma>" by (simp add: subtype_implies_closed) | |
| 1023 | with `\<turnstile> (\<Delta>@[(TVarB X Q)]@\<Gamma>) ok` have "\<turnstile> (\<Delta>@[(TVarB X P)]@\<Gamma>) ok" | |
| 1024 | by (simp add: replace_type) | |
| 18424 | 1025 | moreover | 
| 34915 | 1026 | from `Y \<in> ty_dom (\<Delta>@[(TVarB X Q)]@\<Gamma>)` have "Y \<in> ty_dom (\<Delta>@[(TVarB X P)]@\<Gamma>)" | 
| 1027 | by (simp add: doms_append) | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1028 | ultimately show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Tvar Y <: Tvar Y" by (simp add: subtype_of.SA_refl_TVar) | 
| 18246 | 1029 | next | 
| 34915 | 1030 | case (SA_arrow S\<^isub>1 Q\<^isub>1 Q\<^isub>2 S\<^isub>2 \<Gamma> X \<Delta>) | 
| 32011 | 1031 | then show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Q\<^isub>1 \<rightarrow> Q\<^isub>2 <: S\<^isub>1 \<rightarrow> S\<^isub>2" by blast | 
| 18424 | 1032 | next | 
| 34915 | 1033 | case (SA_all T\<^isub>1 S\<^isub>1 Y S\<^isub>2 T\<^isub>2 \<Gamma> X \<Delta>) | 
| 32011 | 1034 | have IH_inner\<^isub>1: "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> T\<^isub>1 <: S\<^isub>1" | 
| 34915 | 1035 | and IH_inner\<^isub>2: "(((TVarB Y T\<^isub>1)#\<Delta>)@[(TVarB X P)]@\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2" | 
| 1036 | by (fastsimp intro: SA_all)+ | |
| 32011 | 1037 | then show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> (\<forall>Y<:S\<^isub>1. S\<^isub>2) <: (\<forall>Y<:T\<^isub>1. T\<^isub>2)" by auto | 
| 18246 | 1038 | qed | 
| 18621 | 1039 | } | 
| 18246 | 1040 | qed | 
| 1041 | ||
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1042 | section {* Typing *}
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1043 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1044 | inductive | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1045 |   typing :: "env \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _" [60,60,60] 60) 
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1046 | where | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1047 | T_Var[intro]: "\<lbrakk> VarB x T \<in> set \<Gamma>; \<turnstile> \<Gamma> ok \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1048 | | T_App[intro]: "\<lbrakk> \<Gamma> \<turnstile> t\<^isub>1 : T\<^isub>1 \<rightarrow> T\<^isub>2; \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>1 \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> t\<^isub>1 \<cdot> t\<^isub>2 : T\<^isub>2" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1049 | | T_Abs[intro]: "\<lbrakk> VarB x T\<^isub>1 # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2 \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (\<lambda>x:T\<^isub>1. t\<^isub>2) : T\<^isub>1 \<rightarrow> T\<^isub>2" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1050 | | T_Sub[intro]: "\<lbrakk> \<Gamma> \<turnstile> t : S; \<Gamma> \<turnstile> S <: T \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> t : T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1051 | | T_TAbs[intro]:"\<lbrakk> TVarB X T\<^isub>1 # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2 \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (\<lambda>X<:T\<^isub>1. t\<^isub>2) : (\<forall>X<:T\<^isub>1. T\<^isub>2)" | 
| 35438 | 1052 | | T_TApp[intro]:"\<lbrakk>X\<sharp>(\<Gamma>,t\<^isub>1,T\<^isub>2); \<Gamma> \<turnstile> t\<^isub>1 : (\<forall>X<:T\<^isub>1\<^isub>1. T\<^isub>1\<^isub>2); \<Gamma> \<turnstile> T\<^isub>2 <: T\<^isub>1\<^isub>1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> t\<^isub>1 \<cdot>\<^sub>\<tau> T\<^isub>2 : (T\<^isub>1\<^isub>2[X \<mapsto> T\<^isub>2]\<^sub>\<tau>)" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1053 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1054 | equivariance typing | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1055 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1056 | lemma better_T_TApp: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1057 | assumes H1: "\<Gamma> \<turnstile> t\<^isub>1 : (\<forall>X<:T11. T12)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1058 | and H2: "\<Gamma> \<turnstile> T2 <: T11" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1059 | shows "\<Gamma> \<turnstile> t\<^isub>1 \<cdot>\<^sub>\<tau> T2 : (T12[X \<mapsto> T2]\<^sub>\<tau>)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1060 | proof - | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1061 | obtain Y::tyvrs where Y: "Y \<sharp> (X, T12, \<Gamma>, t\<^isub>1, T2)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1062 | by (rule exists_fresh) (rule fin_supp) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1063 | then have "Y \<sharp> (\<Gamma>, t\<^isub>1, T2)" by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1064 | moreover from Y have "(\<forall>X<:T11. T12) = (\<forall>Y<:T11. [(Y, X)] \<bullet> T12)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1065 | by (auto simp add: ty.inject alpha' fresh_prod fresh_atm) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1066 | with H1 have "\<Gamma> \<turnstile> t\<^isub>1 : (\<forall>Y<:T11. [(Y, X)] \<bullet> T12)" by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1067 | ultimately have "\<Gamma> \<turnstile> t\<^isub>1 \<cdot>\<^sub>\<tau> T2 : (([(Y, X)] \<bullet> T12)[Y \<mapsto> T2]\<^sub>\<tau>)" using H2 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1068 | by (rule T_TApp) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1069 | with Y show ?thesis by (simp add: type_subst_rename) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1070 | qed | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1071 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1072 | lemma typing_ok: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1073 | assumes "\<Gamma> \<turnstile> t : T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1074 | shows "\<turnstile> \<Gamma> ok" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1075 | using assms by (induct, auto) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1076 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1077 | nominal_inductive typing | 
| 32011 | 1078 | by (auto dest!: typing_ok intro: closed_in_fresh fresh_dom type_subst_fresh | 
| 1079 | simp: abs_fresh fresh_type_subst_fresh ty_vrs_fresh valid_ty_dom_fresh fresh_trm_dom) | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1080 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1081 | lemma ok_imp_VarB_closed_in: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1082 | assumes ok: "\<turnstile> \<Gamma> ok" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1083 | shows "VarB x T \<in> set \<Gamma> \<Longrightarrow> T closed_in \<Gamma>" using ok | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1084 | by induct (auto simp add: binding.inject closed_in_def) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1085 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1086 | lemma tyvrs_of_subst: "tyvrs_of (B[X \<mapsto> T]\<^sub>b) = tyvrs_of B" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1087 | by (nominal_induct B rule: binding.strong_induct) simp_all | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1088 | |
| 32011 | 1089 | lemma ty_dom_subst: "ty_dom (\<Gamma>[X \<mapsto> T]\<^sub>e) = ty_dom \<Gamma>" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1090 | by (induct \<Gamma>) (simp_all add: tyvrs_of_subst) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1091 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1092 | lemma vrs_of_subst: "vrs_of (B[X \<mapsto> T]\<^sub>b) = vrs_of B" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1093 | by (nominal_induct B rule: binding.strong_induct) simp_all | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1094 | |
| 32011 | 1095 | lemma trm_dom_subst: "trm_dom (\<Gamma>[X \<mapsto> T]\<^sub>e) = trm_dom \<Gamma>" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1096 | by (induct \<Gamma>) (simp_all add: vrs_of_subst) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1097 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1098 | lemma subst_closed_in: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1099 | "T closed_in (\<Delta> @ TVarB X S # \<Gamma>) \<Longrightarrow> U closed_in \<Gamma> \<Longrightarrow> T[X \<mapsto> U]\<^sub>\<tau> closed_in (\<Delta>[X \<mapsto> U]\<^sub>e @ \<Gamma>)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1100 | apply (nominal_induct T avoiding: X U \<Gamma> rule: ty.strong_induct) | 
| 32011 | 1101 | apply (simp add: closed_in_def ty.supp supp_atm doms_append ty_dom_subst) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1102 | apply blast | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1103 | apply (simp add: closed_in_def ty.supp) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1104 | apply (simp add: closed_in_def ty.supp) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1105 | apply (simp add: closed_in_def ty.supp abs_supp) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1106 | apply (drule_tac x = X in meta_spec) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1107 | apply (drule_tac x = U in meta_spec) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1108 | apply (drule_tac x = "(TVarB tyvrs ty2) # \<Gamma>" in meta_spec) | 
| 32011 | 1109 | apply (simp add: doms_append ty_dom_subst) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1110 | apply blast | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1111 | done | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1112 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1113 | lemmas subst_closed_in' = subst_closed_in [where \<Delta>="[]", simplified] | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1114 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1115 | lemma typing_closed_in: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1116 | assumes "\<Gamma> \<turnstile> t : T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1117 | shows "T closed_in \<Gamma>" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1118 | using assms | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1119 | proof induct | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1120 | case (T_Var x T \<Gamma>) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1121 | from `\<turnstile> \<Gamma> ok` and `VarB x T \<in> set \<Gamma>` | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1122 | show ?case by (rule ok_imp_VarB_closed_in) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1123 | next | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1124 | case (T_App \<Gamma> t\<^isub>1 T\<^isub>1 T\<^isub>2 t\<^isub>2) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1125 | then show ?case by (auto simp add: ty.supp closed_in_def) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1126 | next | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1127 | case (T_Abs x T\<^isub>1 \<Gamma> t\<^isub>2 T\<^isub>2) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1128 | from `VarB x T\<^isub>1 # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2` | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1129 | have "T\<^isub>1 closed_in \<Gamma>" by (auto dest: typing_ok) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1130 | with T_Abs show ?case by (auto simp add: ty.supp closed_in_def) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1131 | next | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1132 | case (T_Sub \<Gamma> t S T) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1133 | from `\<Gamma> \<turnstile> S <: T` show ?case by (simp add: subtype_implies_closed) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1134 | next | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1135 | case (T_TAbs X T\<^isub>1 \<Gamma> t\<^isub>2 T\<^isub>2) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1136 | from `TVarB X T\<^isub>1 # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2` | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1137 | have "T\<^isub>1 closed_in \<Gamma>" by (auto dest: typing_ok) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1138 | with T_TAbs show ?case by (auto simp add: ty.supp closed_in_def abs_supp) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1139 | next | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1140 | case (T_TApp X \<Gamma> t\<^isub>1 T2 T11 T12) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1141 | then have "T12 closed_in (TVarB X T11 # \<Gamma>)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1142 | by (auto simp add: closed_in_def ty.supp abs_supp) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1143 | moreover from T_TApp have "T2 closed_in \<Gamma>" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1144 | by (simp add: subtype_implies_closed) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1145 | ultimately show ?case by (rule subst_closed_in') | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1146 | qed | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1147 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1148 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1149 | subsection {* Evaluation *}
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1150 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1151 | inductive | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1152 | val :: "trm \<Rightarrow> bool" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1153 | where | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1154 | Abs[intro]: "val (\<lambda>x:T. t)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1155 | | TAbs[intro]: "val (\<lambda>X<:T. t)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1156 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1157 | equivariance val | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1158 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1159 | inductive_cases val_inv_auto[elim]: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1160 | "val (Var x)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1161 | "val (t1 \<cdot> t2)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1162 | "val (t1 \<cdot>\<^sub>\<tau> t2)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1163 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1164 | inductive | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1165 |   eval :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longmapsto> _" [60,60] 60)
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1166 | where | 
| 35438 | 1167 | E_Abs : "\<lbrakk> x \<sharp> v\<^isub>2; val v\<^isub>2 \<rbrakk> \<Longrightarrow> (\<lambda>x:T\<^isub>1\<^isub>1. t\<^isub>1\<^isub>2) \<cdot> v\<^isub>2 \<longmapsto> t\<^isub>1\<^isub>2[x \<mapsto> v\<^isub>2]" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1168 | | E_App1 [intro]: "t \<longmapsto> t' \<Longrightarrow> t \<cdot> u \<longmapsto> t' \<cdot> u" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1169 | | E_App2 [intro]: "\<lbrakk> val v; t \<longmapsto> t' \<rbrakk> \<Longrightarrow> v \<cdot> t \<longmapsto> v \<cdot> t'" | 
| 35438 | 1170 | | E_TAbs : "X \<sharp> (T\<^isub>1\<^isub>1, T\<^isub>2) \<Longrightarrow> (\<lambda>X<:T\<^isub>1\<^isub>1. t\<^isub>1\<^isub>2) \<cdot>\<^sub>\<tau> T\<^isub>2 \<longmapsto> t\<^isub>1\<^isub>2[X \<mapsto>\<^sub>\<tau> T\<^isub>2]" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1171 | | E_TApp [intro]: "t \<longmapsto> t' \<Longrightarrow> t \<cdot>\<^sub>\<tau> T \<longmapsto> t' \<cdot>\<^sub>\<tau> T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1172 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1173 | lemma better_E_Abs[intro]: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1174 | assumes H: "val v2" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1175 | shows "(\<lambda>x:T11. t12) \<cdot> v2 \<longmapsto> t12[x \<mapsto> v2]" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1176 | proof - | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1177 | obtain y::vrs where y: "y \<sharp> (x, t12, v2)" by (rule exists_fresh) (rule fin_supp) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1178 | then have "y \<sharp> v2" by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1179 | then have "(\<lambda>y:T11. [(y, x)] \<bullet> t12) \<cdot> v2 \<longmapsto> ([(y, x)] \<bullet> t12)[y \<mapsto> v2]" using H | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1180 | by (rule E_Abs) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1181 | moreover from y have "(\<lambda>x:T11. t12) \<cdot> v2 = (\<lambda>y:T11. [(y, x)] \<bullet> t12) \<cdot> v2" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1182 | by (auto simp add: trm.inject alpha' fresh_prod fresh_atm) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1183 | ultimately have "(\<lambda>x:T11. t12) \<cdot> v2 \<longmapsto> ([(y, x)] \<bullet> t12)[y \<mapsto> v2]" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1184 | by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1185 | with y show ?thesis by (simp add: subst_trm_rename) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1186 | qed | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1187 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1188 | lemma better_E_TAbs[intro]: "(\<lambda>X<:T11. t12) \<cdot>\<^sub>\<tau> T2 \<longmapsto> t12[X \<mapsto>\<^sub>\<tau> T2]" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1189 | proof - | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1190 | obtain Y::tyvrs where Y: "Y \<sharp> (X, t12, T11, T2)" by (rule exists_fresh) (rule fin_supp) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1191 | then have "Y \<sharp> (T11, T2)" by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1192 | then have "(\<lambda>Y<:T11. [(Y, X)] \<bullet> t12) \<cdot>\<^sub>\<tau> T2 \<longmapsto> ([(Y, X)] \<bullet> t12)[Y \<mapsto>\<^sub>\<tau> T2]" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1193 | by (rule E_TAbs) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1194 | moreover from Y have "(\<lambda>X<:T11. t12) \<cdot>\<^sub>\<tau> T2 = (\<lambda>Y<:T11. [(Y, X)] \<bullet> t12) \<cdot>\<^sub>\<tau> T2" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1195 | by (auto simp add: trm.inject alpha' fresh_prod fresh_atm) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1196 | ultimately have "(\<lambda>X<:T11. t12) \<cdot>\<^sub>\<tau> T2 \<longmapsto> ([(Y, X)] \<bullet> t12)[Y \<mapsto>\<^sub>\<tau> T2]" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1197 | by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1198 | with Y show ?thesis by (simp add: subst_trm_ty_rename) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1199 | qed | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1200 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1201 | equivariance eval | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1202 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1203 | nominal_inductive eval | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1204 | by (simp_all add: abs_fresh ty_vrs_fresh subst_trm_fresh_tyvar | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1205 | subst_trm_fresh_var subst_trm_ty_fresh') | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1206 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1207 | inductive_cases eval_inv_auto[elim]: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1208 | "Var x \<longmapsto> t'" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1209 | "(\<lambda>x:T. t) \<longmapsto> t'" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1210 | "(\<lambda>X<:T. t) \<longmapsto> t'" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1211 | |
| 32011 | 1212 | lemma ty_dom_cons: | 
| 1213 | shows "ty_dom (\<Gamma>@[VarB X Q]@\<Delta>) = ty_dom (\<Gamma>@\<Delta>)" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1214 | by (induct \<Gamma>, auto) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1215 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1216 | lemma closed_in_cons: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1217 | assumes "S closed_in (\<Gamma> @ VarB X Q # \<Delta>)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1218 | shows "S closed_in (\<Gamma>@\<Delta>)" | 
| 32011 | 1219 | using assms ty_dom_cons closed_in_def by auto | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1220 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1221 | lemma closed_in_weaken: "T closed_in (\<Delta> @ \<Gamma>) \<Longrightarrow> T closed_in (\<Delta> @ B # \<Gamma>)" | 
| 32011 | 1222 | by (auto simp add: closed_in_def doms_append) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1223 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1224 | lemma closed_in_weaken': "T closed_in \<Gamma> \<Longrightarrow> T closed_in (\<Delta> @ \<Gamma>)" | 
| 32011 | 1225 | by (auto simp add: closed_in_def doms_append) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1226 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1227 | lemma valid_subst: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1228 | assumes ok: "\<turnstile> (\<Delta> @ TVarB X Q # \<Gamma>) ok" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1229 | and closed: "P closed_in \<Gamma>" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1230 | shows "\<turnstile> (\<Delta>[X \<mapsto> P]\<^sub>e @ \<Gamma>) ok" using ok closed | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1231 | apply (induct \<Delta>) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1232 | apply simp_all | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1233 | apply (erule validE) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1234 | apply assumption | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1235 | apply (erule validE) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1236 | apply simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1237 | apply (rule valid_consT) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1238 | apply assumption | 
| 32011 | 1239 | apply (simp add: doms_append ty_dom_subst) | 
| 1240 | apply (simp add: fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_doms) | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1241 | apply (rule_tac S=Q in subst_closed_in') | 
| 32011 | 1242 | apply (simp add: closed_in_def doms_append ty_dom_subst) | 
| 1243 | apply (simp add: closed_in_def doms_append) | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1244 | apply blast | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1245 | apply simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1246 | apply (rule valid_cons) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1247 | apply assumption | 
| 32011 | 1248 | apply (simp add: doms_append trm_dom_subst) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1249 | apply (rule_tac S=Q in subst_closed_in') | 
| 32011 | 1250 | apply (simp add: closed_in_def doms_append ty_dom_subst) | 
| 1251 | apply (simp add: closed_in_def doms_append) | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1252 | apply blast | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1253 | done | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1254 | |
| 32011 | 1255 | lemma ty_dom_vrs: | 
| 1256 | shows "ty_dom (G @ [VarB x Q] @ D) = ty_dom (G @ D)" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1257 | by (induct G, auto) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1258 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1259 | lemma valid_cons': | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1260 | assumes "\<turnstile> (\<Gamma> @ VarB x Q # \<Delta>) ok" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1261 | shows "\<turnstile> (\<Gamma> @ \<Delta>) ok" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1262 | using assms | 
| 34915 | 1263 | proof (induct "\<Gamma> @ VarB x Q # \<Delta>" arbitrary: \<Gamma> \<Delta>) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1264 | case valid_nil | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1265 | have "[] = \<Gamma> @ VarB x Q # \<Delta>" by fact | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1266 | then have "False" by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1267 | then show ?case by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1268 | next | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1269 | case (valid_consT G X T) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1270 | then show ?case | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1271 | proof (cases \<Gamma>) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1272 | case Nil | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1273 | with valid_consT show ?thesis by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1274 | next | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1275 | case (Cons b bs) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1276 | with valid_consT | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1277 | have "\<turnstile> (bs @ \<Delta>) ok" by simp | 
| 32011 | 1278 | moreover from Cons and valid_consT have "X \<sharp> ty_dom (bs @ \<Delta>)" | 
| 1279 | by (simp add: doms_append) | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1280 | moreover from Cons and valid_consT have "T closed_in (bs @ \<Delta>)" | 
| 32011 | 1281 | by (simp add: closed_in_def doms_append) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1282 | ultimately have "\<turnstile> (TVarB X T # bs @ \<Delta>) ok" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1283 | by (rule valid_rel.valid_consT) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1284 | with Cons and valid_consT show ?thesis by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1285 | qed | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1286 | next | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1287 | case (valid_cons G x T) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1288 | then show ?case | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1289 | proof (cases \<Gamma>) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1290 | case Nil | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1291 | with valid_cons show ?thesis by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1292 | next | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1293 | case (Cons b bs) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1294 | with valid_cons | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1295 | have "\<turnstile> (bs @ \<Delta>) ok" by simp | 
| 32011 | 1296 | moreover from Cons and valid_cons have "x \<sharp> trm_dom (bs @ \<Delta>)" | 
| 1297 | by (simp add: doms_append finite_doms | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 1298 | fresh_fin_insert [OF pt_vrs_inst at_vrs_inst fs_vrs_inst]) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1299 | moreover from Cons and valid_cons have "T closed_in (bs @ \<Delta>)" | 
| 32011 | 1300 | by (simp add: closed_in_def doms_append) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1301 | ultimately have "\<turnstile> (VarB x T # bs @ \<Delta>) ok" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1302 | by (rule valid_rel.valid_cons) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1303 | with Cons and valid_cons show ?thesis by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1304 | qed | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1305 | qed | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1306 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1307 | text {* A.5(6) *}
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1308 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1309 | lemma type_weaken: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1310 | assumes "(\<Delta>@\<Gamma>) \<turnstile> t : T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1311 | and "\<turnstile> (\<Delta> @ B # \<Gamma>) ok" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1312 | shows "(\<Delta> @ B # \<Gamma>) \<turnstile> t : T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1313 | using assms | 
| 34915 | 1314 | proof(nominal_induct "\<Delta> @ \<Gamma>" t T avoiding: \<Delta> \<Gamma> B rule: typing.strong_induct) | 
| 1315 | case (T_Var x T) | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1316 | then show ?case by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1317 | next | 
| 35438 | 1318 | case (T_App X t\<^isub>1 T\<^isub>2 T\<^isub>1\<^isub>1 T\<^isub>1\<^isub>2) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1319 | then show ?case by force | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1320 | next | 
| 34915 | 1321 | case (T_Abs y T\<^isub>1 t\<^isub>2 T\<^isub>2 \<Delta> \<Gamma>) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1322 | then have "VarB y T\<^isub>1 # \<Delta> @ \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2" by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1323 | then have closed: "T\<^isub>1 closed_in (\<Delta> @ \<Gamma>)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1324 | by (auto dest: typing_ok) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1325 | have "\<turnstile> (VarB y T\<^isub>1 # \<Delta> @ B # \<Gamma>) ok" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1326 | apply (rule valid_cons) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1327 | apply (rule T_Abs) | 
| 32011 | 1328 | apply (simp add: doms_append | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1329 | fresh_fin_insert [OF pt_vrs_inst at_vrs_inst fs_vrs_inst] | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1330 | fresh_fin_union [OF pt_vrs_inst at_vrs_inst fs_vrs_inst] | 
| 32011 | 1331 | finite_doms finite_vrs fresh_vrs_of T_Abs fresh_trm_dom) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1332 | apply (rule closed_in_weaken) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1333 | apply (rule closed) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1334 | done | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1335 | then have "\<turnstile> ((VarB y T\<^isub>1 # \<Delta>) @ B # \<Gamma>) ok" by simp | 
| 34915 | 1336 | with _ have "(VarB y T\<^isub>1 # \<Delta>) @ B # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2" | 
| 1337 | by (rule T_Abs) simp | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1338 | then have "VarB y T\<^isub>1 # \<Delta> @ B # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2" by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1339 | then show ?case by (rule typing.T_Abs) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1340 | next | 
| 34915 | 1341 | case (T_Sub t S T \<Delta> \<Gamma>) | 
| 1342 | from refl and `\<turnstile> (\<Delta> @ B # \<Gamma>) ok` | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1343 | have "\<Delta> @ B # \<Gamma> \<turnstile> t : S" by (rule T_Sub) | 
| 34915 | 1344 | moreover from `(\<Delta> @ \<Gamma>)\<turnstile>S<:T` and `\<turnstile> (\<Delta> @ B # \<Gamma>) ok` | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1345 | have "(\<Delta> @ B # \<Gamma>)\<turnstile>S<:T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1346 | by (rule weakening) (simp add: extends_def T_Sub) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1347 | ultimately show ?case by (rule typing.T_Sub) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1348 | next | 
| 34915 | 1349 | case (T_TAbs X T\<^isub>1 t\<^isub>2 T\<^isub>2 \<Delta> \<Gamma>) | 
| 1350 | from `TVarB X T\<^isub>1 # \<Delta> @ \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2` | |
| 1351 | have closed: "T\<^isub>1 closed_in (\<Delta> @ \<Gamma>)" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1352 | by (auto dest: typing_ok) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1353 | have "\<turnstile> (TVarB X T\<^isub>1 # \<Delta> @ B # \<Gamma>) ok" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1354 | apply (rule valid_consT) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1355 | apply (rule T_TAbs) | 
| 32011 | 1356 | apply (simp add: doms_append | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1357 | fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1358 | fresh_fin_union [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] | 
| 32011 | 1359 | finite_doms finite_vrs tyvrs_fresh T_TAbs fresh_dom) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1360 | apply (rule closed_in_weaken) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1361 | apply (rule closed) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1362 | done | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1363 | then have "\<turnstile> ((TVarB X T\<^isub>1 # \<Delta>) @ B # \<Gamma>) ok" by simp | 
| 34915 | 1364 | with _ have "(TVarB X T\<^isub>1 # \<Delta>) @ B # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2" | 
| 1365 | by (rule T_TAbs) simp | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1366 | then have "TVarB X T\<^isub>1 # \<Delta> @ B # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2" by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1367 | then show ?case by (rule typing.T_TAbs) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1368 | next | 
| 34915 | 1369 | case (T_TApp X t\<^isub>1 T2 T11 T12 \<Delta> \<Gamma>) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1370 | have "\<Delta> @ B # \<Gamma> \<turnstile> t\<^isub>1 : (\<forall>X<:T11. T12)" | 
| 34915 | 1371 | by (rule T_TApp refl)+ | 
| 1372 | moreover from `(\<Delta> @ \<Gamma>)\<turnstile>T2<:T11` and `\<turnstile> (\<Delta> @ B # \<Gamma>) ok` | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1373 | have "(\<Delta> @ B # \<Gamma>)\<turnstile>T2<:T11" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1374 | by (rule weakening) (simp add: extends_def T_TApp) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1375 | ultimately show ?case by (rule better_T_TApp) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1376 | qed | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1377 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1378 | lemma type_weaken': | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1379 | "\<Gamma> \<turnstile> t : T \<Longrightarrow> \<turnstile> (\<Delta>@\<Gamma>) ok \<Longrightarrow> (\<Delta>@\<Gamma>) \<turnstile> t : T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1380 | apply (induct \<Delta>) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1381 | apply simp_all | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1382 | apply (erule validE) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1383 | apply (insert type_weaken [of "[]", simplified]) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1384 | apply simp_all | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1385 | done | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1386 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1387 | text {* A.6 *}
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1388 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1389 | lemma strengthening: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1390 | assumes "(\<Gamma> @ VarB x Q # \<Delta>) \<turnstile> S <: T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1391 | shows "(\<Gamma>@\<Delta>) \<turnstile> S <: T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1392 | using assms | 
| 34915 | 1393 | proof (induct "\<Gamma> @ VarB x Q # \<Delta>" S T arbitrary: \<Gamma>) | 
| 1394 | case (SA_Top S) | |
| 1395 | then have "\<turnstile> (\<Gamma> @ \<Delta>) ok" by (auto dest: valid_cons') | |
| 1396 | moreover have "S closed_in (\<Gamma> @ \<Delta>)" using SA_Top by (auto dest: closed_in_cons) | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1397 | ultimately show ?case using subtype_of.SA_Top by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1398 | next | 
| 34915 | 1399 | case (SA_refl_TVar X) | 
| 1400 | from `\<turnstile> (\<Gamma> @ VarB x Q # \<Delta>) ok` | |
| 1401 | have h1:"\<turnstile> (\<Gamma> @ \<Delta>) ok" by (auto dest: valid_cons') | |
| 1402 | have "X \<in> ty_dom (\<Gamma> @ VarB x Q # \<Delta>)" using SA_refl_TVar by auto | |
| 1403 | then have h2:"X \<in> ty_dom (\<Gamma> @ \<Delta>)" using ty_dom_vrs by auto | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1404 | show ?case using h1 h2 by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1405 | next | 
| 34915 | 1406 | case (SA_all T1 S1 X S2 T2) | 
| 1407 | have h1:"((TVarB X T1 # \<Gamma>) @ \<Delta>)\<turnstile>S2<:T2" by (fastsimp intro: SA_all) | |
| 1408 | have h2:"(\<Gamma> @ \<Delta>)\<turnstile>T1<:S1" using SA_all by auto | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1409 | then show ?case using h1 h2 by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1410 | qed (auto) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1411 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1412 | lemma narrow_type: -- {* A.7 *}
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1413 | assumes H: "\<Delta> @ (TVarB X Q) # \<Gamma> \<turnstile> t : T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1414 | shows "\<Gamma> \<turnstile> P <: Q \<Longrightarrow> \<Delta> @ (TVarB X P) # \<Gamma> \<turnstile> t : T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1415 | using H | 
| 34915 | 1416 | proof (nominal_induct "\<Delta> @ (TVarB X Q) # \<Gamma>" t T avoiding: P arbitrary: \<Delta> rule: typing.strong_induct) | 
| 1417 | case (T_Var x T P D) | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1418 | then have "VarB x T \<in> set (D @ TVarB X P # \<Gamma>)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1419 | and "\<turnstile> (D @ TVarB X P # \<Gamma>) ok" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1420 | by (auto intro: replace_type dest!: subtype_implies_closed) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1421 | then show ?case by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1422 | next | 
| 34915 | 1423 | case (T_App t1 T1 T2 t2 P D) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1424 | then show ?case by force | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1425 | next | 
| 34915 | 1426 | case (T_Abs x T1 t2 T2 P D) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1427 | then show ?case by (fastsimp dest: typing_ok) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1428 | next | 
| 34915 | 1429 | case (T_Sub t S T P D) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1430 | then show ?case using subtype_narrow by fastsimp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1431 | next | 
| 34915 | 1432 | case (T_TAbs X' T1 t2 T2 P D) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1433 | then show ?case by (fastsimp dest: typing_ok) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1434 | next | 
| 34915 | 1435 | case (T_TApp X' t1 T2 T11 T12 P D) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1436 | then have "D @ TVarB X P # \<Gamma> \<turnstile> t1 : Forall X' T12 T11" by fastsimp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1437 | moreover have "(D @ [TVarB X Q] @ \<Gamma>) \<turnstile> T2<:T11" using T_TApp by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1438 | then have "(D @ [TVarB X P] @ \<Gamma>) \<turnstile> T2<:T11" using `\<Gamma>\<turnstile>P<:Q` | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1439 | by (rule subtype_narrow) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1440 | moreover from T_TApp have "X' \<sharp> (D @ TVarB X P # \<Gamma>, t1, T2)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1441 | by (simp add: fresh_list_append fresh_list_cons fresh_prod) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1442 | ultimately show ?case by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1443 | qed | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1444 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1445 | subsection {* Substitution lemmas *}
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1446 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1447 | subsubsection {* Substition Preserves Typing *}
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1448 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1449 | theorem subst_type: -- {* A.8 *}
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1450 | assumes H: "(\<Delta> @ (VarB x U) # \<Gamma>) \<turnstile> t : T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1451 | shows "\<Gamma> \<turnstile> u : U \<Longrightarrow> \<Delta> @ \<Gamma> \<turnstile> t[x \<mapsto> u] : T" using H | 
| 34915 | 1452 | proof (nominal_induct "\<Delta> @ (VarB x U) # \<Gamma>" t T avoiding: x u arbitrary: \<Delta> rule: typing.strong_induct) | 
| 1453 | case (T_Var y T x u D) | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1454 | show ?case | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1455 | proof (cases "x = y") | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1456 | assume eq:"x=y" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1457 | then have "T=U" using T_Var uniqueness_of_ctxt' by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1458 | then show ?case using eq T_Var | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1459 | by (auto intro: type_weaken' dest: valid_cons') | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1460 | next | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1461 | assume "x\<noteq>y" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1462 | then show ?case using T_Var | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1463 | by (auto simp add:binding.inject dest: valid_cons') | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1464 | qed | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1465 | next | 
| 34915 | 1466 | case (T_App t1 T1 T2 t2 x u D) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1467 | then show ?case by force | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1468 | next | 
| 34915 | 1469 | case (T_Abs y T1 t2 T2 x u D) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1470 | then show ?case by force | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1471 | next | 
| 34915 | 1472 | case (T_Sub t S T x u D) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1473 | then have "D @ \<Gamma> \<turnstile> t[x \<mapsto> u] : S" by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1474 | moreover have "(D @ \<Gamma>) \<turnstile> S<:T" using T_Sub by (auto dest: strengthening) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1475 | ultimately show ?case by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1476 | next | 
| 34915 | 1477 | case (T_TAbs X T1 t2 T2 x u D) | 
| 1478 | from `TVarB X T1 # D @ VarB x U # \<Gamma> \<turnstile> t2 : T2` have "X \<sharp> T1" | |
| 32011 | 1479 | by (auto simp add: valid_ty_dom_fresh dest: typing_ok intro!: closed_in_fresh) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1480 | with `X \<sharp> u` and T_TAbs show ?case by fastsimp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1481 | next | 
| 34915 | 1482 | case (T_TApp X t1 T2 T11 T12 x u D) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1483 | then have "(D@\<Gamma>) \<turnstile>T2<:T11" using T_TApp by (auto dest: strengthening) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1484 | then show "((D @ \<Gamma>) \<turnstile> ((t1 \<cdot>\<^sub>\<tau> T2)[x \<mapsto> u]) : (T12[X \<mapsto> T2]\<^sub>\<tau>))" using T_TApp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1485 | by (force simp add: fresh_prod fresh_list_append fresh_list_cons subst_trm_fresh_tyvar) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1486 | qed | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1487 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1488 | subsubsection {* Type Substitution Preserves Subtyping *}
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1489 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1490 | lemma substT_subtype: -- {* A.10 *}
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1491 | assumes H: "(\<Delta> @ ((TVarB X Q) # \<Gamma>)) \<turnstile> S <: T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1492 | shows "\<Gamma> \<turnstile> P <: Q \<Longrightarrow> (\<Delta>[X \<mapsto> P]\<^sub>e @ \<Gamma>) \<turnstile> S[X \<mapsto> P]\<^sub>\<tau> <: T[X \<mapsto> P]\<^sub>\<tau>" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1493 | using H | 
| 34915 | 1494 | proof (nominal_induct "\<Delta> @ TVarB X Q # \<Gamma>" S T avoiding: X P arbitrary: \<Delta> rule: subtype_of.strong_induct) | 
| 1495 | case (SA_Top S X P D) | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1496 | then have "\<turnstile> (D @ TVarB X Q # \<Gamma>) ok" by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1497 | moreover have closed: "P closed_in \<Gamma>" using SA_Top subtype_implies_closed by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1498 | ultimately have "\<turnstile> (D[X \<mapsto> P]\<^sub>e @ \<Gamma>) ok" by (rule valid_subst) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1499 | moreover from SA_Top have "S closed_in (D @ TVarB X Q # \<Gamma>)" by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1500 | then have "S[X \<mapsto> P]\<^sub>\<tau> closed_in (D[X \<mapsto> P]\<^sub>e @ \<Gamma>)" using closed by (rule subst_closed_in) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1501 | ultimately show ?case by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1502 | next | 
| 34915 | 1503 | case (SA_trans_TVar Y S T X P D) | 
| 1504 | have h:"(D @ TVarB X Q # \<Gamma>)\<turnstile>S<:T" by fact | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1505 | then have ST: "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>) \<turnstile> S[X \<mapsto> P]\<^sub>\<tau> <: T[X \<mapsto> P]\<^sub>\<tau>" using SA_trans_TVar by auto | 
| 34915 | 1506 | from h have G_ok: "\<turnstile> (D @ TVarB X Q # \<Gamma>) ok" by (rule subtype_implies_ok) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1507 | from G_ok and SA_trans_TVar have X\<Gamma>_ok: "\<turnstile> (TVarB X Q # \<Gamma>) ok" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1508 | by (auto intro: validE_append) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1509 | show "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>) \<turnstile> Tvar Y[X \<mapsto> P]\<^sub>\<tau><:T[X \<mapsto> P]\<^sub>\<tau>" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1510 | proof (cases "X = Y") | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1511 | assume eq: "X = Y" | 
| 34915 | 1512 | from eq and SA_trans_TVar have "TVarB Y Q \<in> set (D @ TVarB X Q # \<Gamma>)" by simp | 
| 1513 | with G_ok have QS: "Q = S" using `TVarB Y S \<in> set (D @ TVarB X Q # \<Gamma>)` | |
| 1514 | by (rule uniqueness_of_ctxt) | |
| 32011 | 1515 | from X\<Gamma>_ok have "X \<sharp> ty_dom \<Gamma>" and "Q closed_in \<Gamma>" by auto | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1516 | then have XQ: "X \<sharp> Q" by (rule closed_in_fresh) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1517 | note `\<Gamma>\<turnstile>P<:Q` | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1518 | moreover from ST have "\<turnstile> (D[X \<mapsto> P]\<^sub>e @ \<Gamma>) ok" by (rule subtype_implies_ok) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1519 | moreover have "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>) extends \<Gamma>" by (simp add: extends_def) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1520 | ultimately have "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>) \<turnstile> P<:Q" by (rule weakening) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1521 | with QS have "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>) \<turnstile> P<:S" by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1522 | moreover from XQ and ST and QS have "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>) \<turnstile> S<:T[X \<mapsto> P]\<^sub>\<tau>" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1523 | by (simp add: type_subst_identity) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1524 | ultimately have "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>) \<turnstile> P<:T[X \<mapsto> P]\<^sub>\<tau>" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1525 | by (rule subtype_transitivity) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1526 | with eq show ?case by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1527 | next | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1528 | assume neq: "X \<noteq> Y" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1529 | with SA_trans_TVar have "TVarB Y S \<in> set D \<or> TVarB Y S \<in> set \<Gamma>" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1530 | by (simp add: binding.inject) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1531 | then show ?case | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1532 | proof | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1533 | assume "TVarB Y S \<in> set D" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1534 | then have "TVarB Y (S[X \<mapsto> P]\<^sub>\<tau>) \<in> set (D[X \<mapsto> P]\<^sub>e)" | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 1535 | by (rule ctxt_subst_mem_TVarB) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1536 | then have "TVarB Y (S[X \<mapsto> P]\<^sub>\<tau>) \<in> set (D[X \<mapsto> P]\<^sub>e @ \<Gamma>)" by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1537 | with neq and ST show ?thesis by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1538 | next | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1539 | assume Y: "TVarB Y S \<in> set \<Gamma>" | 
| 32011 | 1540 | from X\<Gamma>_ok have "X \<sharp> ty_dom \<Gamma>" and "\<turnstile> \<Gamma> ok" by auto | 
| 1541 | then have "X \<sharp> \<Gamma>" by (simp add: valid_ty_dom_fresh) | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1542 | with Y have "X \<sharp> S" | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 1543 | by (induct \<Gamma>) (auto simp add: fresh_list_nil fresh_list_cons) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1544 | with ST have "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>)\<turnstile>S<:T[X \<mapsto> P]\<^sub>\<tau>" | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32011diff
changeset | 1545 | by (simp add: type_subst_identity) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1546 | moreover from Y have "TVarB Y S \<in> set (D[X \<mapsto> P]\<^sub>e @ \<Gamma>)" by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1547 | ultimately show ?thesis using neq by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1548 | qed | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1549 | qed | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1550 | next | 
| 34915 | 1551 | case (SA_refl_TVar Y X P D) | 
| 1552 | note `\<turnstile> (D @ TVarB X Q # \<Gamma>) ok` | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1553 | moreover from SA_refl_TVar have closed: "P closed_in \<Gamma>" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1554 | by (auto dest: subtype_implies_closed) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1555 | ultimately have ok: "\<turnstile> (D[X \<mapsto> P]\<^sub>e @ \<Gamma>) ok" using valid_subst by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1556 | from closed have closed': "P closed_in (D[X \<mapsto> P]\<^sub>e @ \<Gamma>)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1557 | by (simp add: closed_in_weaken') | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1558 | show ?case | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1559 | proof (cases "X = Y") | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1560 | assume "X = Y" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1561 | with closed' and ok show ?thesis | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1562 | by (auto intro: subtype_reflexivity) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1563 | next | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1564 | assume neq: "X \<noteq> Y" | 
| 32011 | 1565 | with SA_refl_TVar have "Y \<in> ty_dom (D[X \<mapsto> P]\<^sub>e @ \<Gamma>)" | 
| 1566 | by (simp add: ty_dom_subst doms_append) | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1567 | with neq and ok show ?thesis by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1568 | qed | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1569 | next | 
| 34915 | 1570 | case (SA_arrow T1 S1 S2 T2 X P D) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1571 | then have h1:"(D[X \<mapsto> P]\<^sub>e @ \<Gamma>)\<turnstile>T1[X \<mapsto> P]\<^sub>\<tau><:S1[X \<mapsto> P]\<^sub>\<tau>" using SA_arrow by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1572 | from SA_arrow have h2:"(D[X \<mapsto> P]\<^sub>e @ \<Gamma>)\<turnstile>S2[X \<mapsto> P]\<^sub>\<tau><:T2[X \<mapsto> P]\<^sub>\<tau>" using SA_arrow by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1573 | show ?case using subtype_of.SA_arrow h1 h2 by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1574 | next | 
| 34915 | 1575 | case (SA_all T1 S1 Y S2 T2 X P D) | 
| 32011 | 1576 | then have Y: "Y \<sharp> ty_dom (D @ TVarB X Q # \<Gamma>)" | 
| 1577 | by (auto dest: subtype_implies_ok intro: fresh_dom) | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1578 | moreover from SA_all have "S1 closed_in (D @ TVarB X Q # \<Gamma>)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1579 | by (auto dest: subtype_implies_closed) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1580 | ultimately have S1: "Y \<sharp> S1" by (rule closed_in_fresh) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1581 | from SA_all have "T1 closed_in (D @ TVarB X Q # \<Gamma>)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1582 | by (auto dest: subtype_implies_closed) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1583 | with Y have T1: "Y \<sharp> T1" by (rule closed_in_fresh) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1584 | with SA_all and S1 show ?case by force | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1585 | qed | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1586 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1587 | subsubsection {* Type Substitution Preserves Typing *}
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1588 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1589 | theorem substT_type: -- {* A.11 *}
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1590 | assumes H: "(D @ TVarB X Q # G) \<turnstile> t : T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1591 | shows "G \<turnstile> P <: Q \<Longrightarrow> | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1592 | (D[X \<mapsto> P]\<^sub>e @ G) \<turnstile> t[X \<mapsto>\<^sub>\<tau> P] : T[X \<mapsto> P]\<^sub>\<tau>" using H | 
| 34915 | 1593 | proof (nominal_induct "D @ TVarB X Q # G" t T avoiding: X P arbitrary: D rule: typing.strong_induct) | 
| 1594 | case (T_Var x T X P D') | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1595 | have "G\<turnstile>P<:Q" by fact | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1596 | then have "P closed_in G" using subtype_implies_closed by auto | 
| 34915 | 1597 | moreover note `\<turnstile> (D' @ TVarB X Q # G) ok` | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1598 | ultimately have "\<turnstile> (D'[X \<mapsto> P]\<^sub>e @ G) ok" using valid_subst by auto | 
| 34915 | 1599 | moreover note `VarB x T \<in> set (D' @ TVarB X Q # G)` | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1600 | then have "VarB x T \<in> set D' \<or> VarB x T \<in> set G" by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1601 | then have "(VarB x (T[X \<mapsto> P]\<^sub>\<tau>)) \<in> set (D'[X \<mapsto> P]\<^sub>e @ G)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1602 | proof | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1603 | assume "VarB x T \<in> set D'" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1604 | then have "VarB x (T[X \<mapsto> P]\<^sub>\<tau>) \<in> set (D'[X \<mapsto> P]\<^sub>e)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1605 | by (rule ctxt_subst_mem_VarB) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1606 | then show ?thesis by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1607 | next | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1608 | assume x: "VarB x T \<in> set G" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1609 | from T_Var have ok: "\<turnstile> G ok" by (auto dest: subtype_implies_ok) | 
| 32011 | 1610 | then have "X \<sharp> ty_dom G" using T_Var by (auto dest: validE_append) | 
| 1611 | with ok have "X \<sharp> G" by (simp add: valid_ty_dom_fresh) | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1612 | moreover from x have "VarB x T \<in> set (D' @ G)" by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1613 | then have "VarB x (T[X \<mapsto> P]\<^sub>\<tau>) \<in> set ((D' @ G)[X \<mapsto> P]\<^sub>e)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1614 | by (rule ctxt_subst_mem_VarB) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1615 | ultimately show ?thesis | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1616 | by (simp add: ctxt_subst_append ctxt_subst_identity) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1617 | qed | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1618 | ultimately show ?case by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1619 | next | 
| 34915 | 1620 | case (T_App t1 T1 T2 t2 X P D') | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1621 | then have "D'[X \<mapsto> P]\<^sub>e @ G \<turnstile> t1[X \<mapsto>\<^sub>\<tau> P] : (T1 \<rightarrow> T2)[X \<mapsto> P]\<^sub>\<tau>" by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1622 | moreover from T_App have "D'[X \<mapsto> P]\<^sub>e @ G \<turnstile> t2[X \<mapsto>\<^sub>\<tau> P] : T1[X \<mapsto> P]\<^sub>\<tau>" by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1623 | ultimately show ?case by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1624 | next | 
| 34915 | 1625 | case (T_Abs x T1 t2 T2 X P D') | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1626 | then show ?case by force | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1627 | next | 
| 34915 | 1628 | case (T_Sub t S T X P D') | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1629 | then show ?case using substT_subtype by force | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1630 | next | 
| 34915 | 1631 | case (T_TAbs X' T1 t2 T2 X P D') | 
| 32011 | 1632 | then have "X' \<sharp> ty_dom (D' @ TVarB X Q # G)" | 
| 34915 | 1633 | and "T1 closed_in (D' @ TVarB X Q # G)" | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1634 | by (auto dest: typing_ok) | 
| 34915 | 1635 | then have "X' \<sharp> T1" by (rule closed_in_fresh) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1636 | with T_TAbs show ?case by force | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1637 | next | 
| 34915 | 1638 | case (T_TApp X' t1 T2 T11 T12 X P D') | 
| 32011 | 1639 | then have "X' \<sharp> ty_dom (D' @ TVarB X Q # G)" | 
| 1640 | by (simp add: fresh_dom) | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1641 | moreover from T_TApp have "T11 closed_in (D' @ TVarB X Q # G)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1642 | by (auto dest: subtype_implies_closed) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1643 | ultimately have X': "X' \<sharp> T11" by (rule closed_in_fresh) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1644 | from T_TApp have "D'[X \<mapsto> P]\<^sub>e @ G \<turnstile> t1[X \<mapsto>\<^sub>\<tau> P] : (\<forall>X'<:T11. T12)[X \<mapsto> P]\<^sub>\<tau>" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1645 | by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1646 | with X' and T_TApp show ?case | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1647 | by (auto simp add: fresh_atm type_substitution_lemma | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1648 | fresh_list_append fresh_list_cons | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1649 | ctxt_subst_fresh' type_subst_fresh subst_trm_ty_fresh | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1650 | intro: substT_subtype) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1651 | qed | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1652 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1653 | lemma Abs_type: -- {* A.13(1) *}
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1654 | assumes H: "\<Gamma> \<turnstile> (\<lambda>x:S. s) : T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1655 | and H': "\<Gamma> \<turnstile> T <: U \<rightarrow> U'" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1656 | and H'': "x \<sharp> \<Gamma>" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1657 | obtains S' where "\<Gamma> \<turnstile> U <: S" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1658 | and "(VarB x S) # \<Gamma> \<turnstile> s : S'" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1659 | and "\<Gamma> \<turnstile> S' <: U'" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1660 | using H H' H'' | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1661 | proof (nominal_induct \<Gamma> t \<equiv> "\<lambda>x:S. s" T avoiding: x arbitrary: U U' S s rule: typing.strong_induct) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1662 | case (T_Abs y T\<^isub>1 \<Gamma> t\<^isub>2 T\<^isub>2) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1663 | from `\<Gamma> \<turnstile> T\<^isub>1 \<rightarrow> T\<^isub>2 <: U \<rightarrow> U'` | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1664 | obtain ty1: "\<Gamma> \<turnstile> U <: S" and ty2: "\<Gamma> \<turnstile> T\<^isub>2 <: U'" using T_Abs | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1665 | by cases (simp_all add: ty.inject trm.inject alpha fresh_atm) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1666 | from T_Abs have "VarB y S # \<Gamma> \<turnstile> [(y, x)] \<bullet> s : T\<^isub>2" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1667 | by (simp add: trm.inject alpha fresh_atm) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1668 | then have "[(y, x)] \<bullet> (VarB y S # \<Gamma>) \<turnstile> [(y, x)] \<bullet> [(y, x)] \<bullet> s : [(y, x)] \<bullet> T\<^isub>2" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1669 | by (rule typing.eqvt) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1670 | moreover from T_Abs have "y \<sharp> \<Gamma>" | 
| 32011 | 1671 | by (auto dest!: typing_ok simp add: fresh_trm_dom) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1672 | ultimately have "VarB x S # \<Gamma> \<turnstile> s : T\<^isub>2" using T_Abs | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1673 | by (perm_simp add: ty_vrs_prm_simp) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1674 | with ty1 show ?case using ty2 by (rule T_Abs) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1675 | next | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1676 | case (T_Sub \<Gamma> t S T) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1677 | then show ?case using subtype_transitivity by blast | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1678 | qed simp_all | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1679 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1680 | lemma subtype_reflexivity_from_typing: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1681 | assumes "\<Gamma> \<turnstile> t : T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1682 | shows "\<Gamma> \<turnstile> T <: T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1683 | using assms subtype_reflexivity typing_ok typing_closed_in by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1684 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1685 | lemma Abs_type': | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1686 | assumes H: "\<Gamma> \<turnstile> (\<lambda>x:S. s) : U \<rightarrow> U'" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1687 | and H': "x \<sharp> \<Gamma>" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1688 | obtains S' | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1689 | where "\<Gamma> \<turnstile> U <: S" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1690 | and "(VarB x S) # \<Gamma> \<turnstile> s : S'" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1691 | and "\<Gamma> \<turnstile> S' <: U'" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1692 | using H subtype_reflexivity_from_typing [OF H] H' | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1693 | by (rule Abs_type) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1694 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1695 | lemma TAbs_type: -- {* A.13(2) *}
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1696 | assumes H: "\<Gamma> \<turnstile> (\<lambda>X<:S. s) : T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1697 | and H': "\<Gamma> \<turnstile> T <: (\<forall>X<:U. U')" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1698 | and fresh: "X \<sharp> \<Gamma>" "X \<sharp> S" "X \<sharp> U" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1699 | obtains S' | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1700 | where "\<Gamma> \<turnstile> U <: S" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1701 | and "(TVarB X U # \<Gamma>) \<turnstile> s : S'" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1702 | and "(TVarB X U # \<Gamma>) \<turnstile> S' <: U'" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1703 | using H H' fresh | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1704 | proof (nominal_induct \<Gamma> t \<equiv> "\<lambda>X<:S. s" T avoiding: X U U' S arbitrary: s rule: typing.strong_induct) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1705 | case (T_TAbs Y T\<^isub>1 \<Gamma> t\<^isub>2 T\<^isub>2) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1706 | from `TVarB Y T\<^isub>1 # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2` have Y: "Y \<sharp> \<Gamma>" | 
| 32011 | 1707 | by (auto dest!: typing_ok simp add: valid_ty_dom_fresh) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1708 | from `Y \<sharp> U'` and `Y \<sharp> X` | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1709 | have "(\<forall>X<:U. U') = (\<forall>Y<:U. [(Y, X)] \<bullet> U')" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1710 | by (simp add: ty.inject alpha' fresh_atm) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1711 | with T_TAbs have "\<Gamma> \<turnstile> (\<forall>Y<:S. T\<^isub>2) <: (\<forall>Y<:U. [(Y, X)] \<bullet> U')" by (simp add: trm.inject) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1712 | then obtain ty1: "\<Gamma> \<turnstile> U <: S" and ty2: "(TVarB Y U # \<Gamma>) \<turnstile> T\<^isub>2 <: ([(Y, X)] \<bullet> U')" using T_TAbs Y | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1713 | by (cases rule: subtype_of.strong_cases [where X=Y]) (simp_all add: ty.inject alpha abs_fresh) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1714 | note ty1 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1715 | moreover from T_TAbs have "TVarB Y S # \<Gamma> \<turnstile> ([(Y, X)] \<bullet> s) : T\<^isub>2" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1716 | by (simp add: trm.inject alpha fresh_atm) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1717 | then have "[(Y, X)] \<bullet> (TVarB Y S # \<Gamma>) \<turnstile> [(Y, X)] \<bullet> [(Y, X)] \<bullet> s : [(Y, X)] \<bullet> T\<^isub>2" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1718 | by (rule typing.eqvt) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1719 | with `X \<sharp> \<Gamma>` `X \<sharp> S` Y `Y \<sharp> S` have "TVarB X S # \<Gamma> \<turnstile> s : [(Y, X)] \<bullet> T\<^isub>2" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1720 | by perm_simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1721 | then have "TVarB X U # \<Gamma> \<turnstile> s : [(Y, X)] \<bullet> T\<^isub>2" using ty1 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1722 | by (rule narrow_type [of "[]", simplified]) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1723 | moreover from ty2 have "([(Y, X)] \<bullet> (TVarB Y U # \<Gamma>)) \<turnstile> ([(Y, X)] \<bullet> T\<^isub>2) <: ([(Y, X)] \<bullet> [(Y, X)] \<bullet> U')" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1724 | by (rule subtype_of.eqvt) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1725 | with `X \<sharp> \<Gamma>` `X \<sharp> U` Y `Y \<sharp> U` have "(TVarB X U # \<Gamma>) \<turnstile> ([(Y, X)] \<bullet> T\<^isub>2) <: U'" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1726 | by perm_simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1727 | ultimately show ?case by (rule T_TAbs) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1728 | next | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1729 | case (T_Sub \<Gamma> t S T) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1730 | then show ?case using subtype_transitivity by blast | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1731 | qed simp_all | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1732 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1733 | lemma TAbs_type': | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1734 | assumes H: "\<Gamma> \<turnstile> (\<lambda>X<:S. s) : (\<forall>X<:U. U')" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1735 | and fresh: "X \<sharp> \<Gamma>" "X \<sharp> S" "X \<sharp> U" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1736 | obtains S' | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1737 | where "\<Gamma> \<turnstile> U <: S" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1738 | and "(TVarB X U # \<Gamma>) \<turnstile> s : S'" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1739 | and "(TVarB X U # \<Gamma>) \<turnstile> S' <: U'" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1740 | using H subtype_reflexivity_from_typing [OF H] fresh | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1741 | by (rule TAbs_type) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1742 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1743 | theorem preservation: -- {* A.20 *}
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1744 | assumes H: "\<Gamma> \<turnstile> t : T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1745 | shows "t \<longmapsto> t' \<Longrightarrow> \<Gamma> \<turnstile> t' : T" using H | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1746 | proof (nominal_induct avoiding: t' rule: typing.strong_induct) | 
| 35438 | 1747 | case (T_App \<Gamma> t\<^isub>1 T\<^isub>1\<^isub>1 T\<^isub>1\<^isub>2 t\<^isub>2 t') | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1748 | obtain x::vrs where x_fresh: "x \<sharp> (\<Gamma>, t\<^isub>1 \<cdot> t\<^isub>2, t')" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1749 | by (rule exists_fresh) (rule fin_supp) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1750 | obtain X::tyvrs where "X \<sharp> (t\<^isub>1 \<cdot> t\<^isub>2, t')" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1751 | by (rule exists_fresh) (rule fin_supp) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1752 | with `t\<^isub>1 \<cdot> t\<^isub>2 \<longmapsto> t'` show ?case | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1753 | proof (cases rule: eval.strong_cases [where x=x and X=X]) | 
| 35438 | 1754 | case (E_Abs v\<^isub>2 T\<^isub>1\<^isub>1' t\<^isub>1\<^isub>2) | 
| 1755 | with T_App and x_fresh have h: "\<Gamma> \<turnstile> (\<lambda>x:T\<^isub>1\<^isub>1'. t\<^isub>1\<^isub>2) : T\<^isub>1\<^isub>1 \<rightarrow> T\<^isub>1\<^isub>2" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1756 | by (simp add: trm.inject fresh_prod) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1757 | moreover from x_fresh have "x \<sharp> \<Gamma>" by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1758 | ultimately obtain S' | 
| 35438 | 1759 | where T\<^isub>1\<^isub>1: "\<Gamma> \<turnstile> T\<^isub>1\<^isub>1 <: T\<^isub>1\<^isub>1'" | 
| 1760 | and t\<^isub>1\<^isub>2: "(VarB x T\<^isub>1\<^isub>1') # \<Gamma> \<turnstile> t\<^isub>1\<^isub>2 : S'" | |
| 1761 | and S': "\<Gamma> \<turnstile> S' <: T\<^isub>1\<^isub>2" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1762 | by (rule Abs_type') blast | 
| 35438 | 1763 | from `\<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>1\<^isub>1` | 
| 1764 | have "\<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>1\<^isub>1'" using T\<^isub>1\<^isub>1 by (rule T_Sub) | |
| 1765 | with t\<^isub>1\<^isub>2 have "\<Gamma> \<turnstile> t\<^isub>1\<^isub>2[x \<mapsto> t\<^isub>2] : S'" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1766 | by (rule subst_type [where \<Delta>="[]", simplified]) | 
| 35438 | 1767 | hence "\<Gamma> \<turnstile> t\<^isub>1\<^isub>2[x \<mapsto> t\<^isub>2] : T\<^isub>1\<^isub>2" using S' by (rule T_Sub) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1768 | with E_Abs and x_fresh show ?thesis by (simp add: trm.inject fresh_prod) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1769 | next | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1770 | case (E_App1 t''' t'' u) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1771 | hence "t\<^isub>1 \<longmapsto> t''" by (simp add:trm.inject) | 
| 35438 | 1772 | hence "\<Gamma> \<turnstile> t'' : T\<^isub>1\<^isub>1 \<rightarrow> T\<^isub>1\<^isub>2" by (rule T_App) | 
| 1773 | hence "\<Gamma> \<turnstile> t'' \<cdot> t\<^isub>2 : T\<^isub>1\<^isub>2" using `\<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>1\<^isub>1` | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1774 | by (rule typing.T_App) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1775 | with E_App1 show ?thesis by (simp add:trm.inject) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1776 | next | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1777 | case (E_App2 v t''' t'') | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1778 | hence "t\<^isub>2 \<longmapsto> t''" by (simp add:trm.inject) | 
| 35438 | 1779 | hence "\<Gamma> \<turnstile> t'' : T\<^isub>1\<^isub>1" by (rule T_App) | 
| 1780 | with T_App(1) have "\<Gamma> \<turnstile> t\<^isub>1 \<cdot> t'' : T\<^isub>1\<^isub>2" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1781 | by (rule typing.T_App) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1782 | with E_App2 show ?thesis by (simp add:trm.inject) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1783 | qed (simp_all add: fresh_prod) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1784 | next | 
| 35438 | 1785 | case (T_TApp X \<Gamma> t\<^isub>1 T\<^isub>2 T\<^isub>1\<^isub>1 T\<^isub>1\<^isub>2 t') | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1786 | obtain x::vrs where "x \<sharp> (t\<^isub>1 \<cdot>\<^sub>\<tau> T\<^isub>2, t')" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1787 | by (rule exists_fresh) (rule fin_supp) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1788 | with `t\<^isub>1 \<cdot>\<^sub>\<tau> T\<^isub>2 \<longmapsto> t'` | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1789 | show ?case | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1790 | proof (cases rule: eval.strong_cases [where X=X and x=x]) | 
| 35438 | 1791 | case (E_TAbs T\<^isub>1\<^isub>1' T\<^isub>2' t\<^isub>1\<^isub>2) | 
| 1792 | with T_TApp have "\<Gamma> \<turnstile> (\<lambda>X<:T\<^isub>1\<^isub>1'. t\<^isub>1\<^isub>2) : (\<forall>X<:T\<^isub>1\<^isub>1. T\<^isub>1\<^isub>2)" and "X \<sharp> \<Gamma>" and "X \<sharp> T\<^isub>1\<^isub>1'" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1793 | by (simp_all add: trm.inject) | 
| 35438 | 1794 | moreover from `\<Gamma>\<turnstile>T\<^isub>2<:T\<^isub>1\<^isub>1` and `X \<sharp> \<Gamma>` have "X \<sharp> T\<^isub>1\<^isub>1" | 
| 32011 | 1795 | by (blast intro: closed_in_fresh fresh_dom dest: subtype_implies_closed) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1796 | ultimately obtain S' | 
| 35438 | 1797 | where "TVarB X T\<^isub>1\<^isub>1 # \<Gamma> \<turnstile> t\<^isub>1\<^isub>2 : S'" | 
| 1798 | and "(TVarB X T\<^isub>1\<^isub>1 # \<Gamma>) \<turnstile> S' <: T\<^isub>1\<^isub>2" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1799 | by (rule TAbs_type') blast | 
| 35438 | 1800 | hence "TVarB X T\<^isub>1\<^isub>1 # \<Gamma> \<turnstile> t\<^isub>1\<^isub>2 : T\<^isub>1\<^isub>2" by (rule T_Sub) | 
| 1801 | hence "\<Gamma> \<turnstile> t\<^isub>1\<^isub>2[X \<mapsto>\<^sub>\<tau> T\<^isub>2] : T\<^isub>1\<^isub>2[X \<mapsto> T\<^isub>2]\<^sub>\<tau>" using `\<Gamma> \<turnstile> T\<^isub>2 <: T\<^isub>1\<^isub>1` | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1802 | by (rule substT_type [where D="[]", simplified]) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1803 | with T_TApp and E_TAbs show ?thesis by (simp add: trm.inject) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1804 | next | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1805 | case (E_TApp t''' t'' T) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1806 | from E_TApp have "t\<^isub>1 \<longmapsto> t''" by (simp add: trm.inject) | 
| 35438 | 1807 | then have "\<Gamma> \<turnstile> t'' : (\<forall>X<:T\<^isub>1\<^isub>1. T\<^isub>1\<^isub>2)" by (rule T_TApp) | 
| 1808 | then have "\<Gamma> \<turnstile> t'' \<cdot>\<^sub>\<tau> T\<^isub>2 : T\<^isub>1\<^isub>2[X \<mapsto> T\<^isub>2]\<^sub>\<tau>" using `\<Gamma> \<turnstile> T\<^isub>2 <: T\<^isub>1\<^isub>1` | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1809 | by (rule better_T_TApp) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1810 | with E_TApp show ?thesis by (simp add: trm.inject) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1811 | qed (simp_all add: fresh_prod) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1812 | next | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1813 | case (T_Sub \<Gamma> t S T t') | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1814 | have "t \<longmapsto> t'" by fact | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1815 | hence "\<Gamma> \<turnstile> t' : S" by (rule T_Sub) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1816 | moreover have "\<Gamma> \<turnstile> S <: T" by fact | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1817 | ultimately show ?case by (rule typing.T_Sub) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1818 | qed (auto) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1819 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1820 | lemma Fun_canonical: -- {* A.14(1) *}
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1821 | assumes ty: "[] \<turnstile> v : T\<^isub>1 \<rightarrow> T\<^isub>2" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1822 | shows "val v \<Longrightarrow> \<exists>x t S. v = (\<lambda>x:S. t)" using ty | 
| 34915 | 1823 | proof (induct "[]::env" v "T\<^isub>1 \<rightarrow> T\<^isub>2" arbitrary: T\<^isub>1 T\<^isub>2) | 
| 1824 | case (T_Sub t S) | |
| 1825 | from `[] \<turnstile> S <: T\<^isub>1 \<rightarrow> T\<^isub>2` | |
| 1826 | obtain S\<^isub>1 S\<^isub>2 where S: "S = S\<^isub>1 \<rightarrow> S\<^isub>2" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1827 | by cases (auto simp add: T_Sub) | 
| 34915 | 1828 | then show ?case using `val t` by (rule T_Sub) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1829 | qed (auto) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1830 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1831 | lemma TyAll_canonical: -- {* A.14(3) *}
 | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1832 | fixes X::tyvrs | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1833 | assumes ty: "[] \<turnstile> v : (\<forall>X<:T\<^isub>1. T\<^isub>2)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1834 | shows "val v \<Longrightarrow> \<exists>X t S. v = (\<lambda>X<:S. t)" using ty | 
| 34915 | 1835 | proof (induct "[]::env" v "\<forall>X<:T\<^isub>1. T\<^isub>2" arbitrary: X T\<^isub>1 T\<^isub>2) | 
| 1836 | case (T_Sub t S) | |
| 1837 | from `[] \<turnstile> S <: (\<forall>X<:T\<^isub>1. T\<^isub>2)` | |
| 1838 | obtain X S\<^isub>1 S\<^isub>2 where S: "S = (\<forall>X<:S\<^isub>1. S\<^isub>2)" | |
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1839 | by cases (auto simp add: T_Sub) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1840 | then show ?case using T_Sub by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1841 | qed (auto) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1842 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1843 | theorem progress: | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1844 | assumes "[] \<turnstile> t : T" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1845 | shows "val t \<or> (\<exists>t'. t \<longmapsto> t')" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1846 | using assms | 
| 34915 | 1847 | proof (induct "[]::env" t T) | 
| 35438 | 1848 | case (T_App t\<^isub>1 T\<^isub>1\<^isub>1 T\<^isub>1\<^isub>2 t\<^isub>2) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1849 | hence "val t\<^isub>1 \<or> (\<exists>t'. t\<^isub>1 \<longmapsto> t')" by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1850 | thus ?case | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1851 | proof | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1852 | assume t\<^isub>1_val: "val t\<^isub>1" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1853 | with T_App obtain x t3 S where t\<^isub>1: "t\<^isub>1 = (\<lambda>x:S. t3)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1854 | by (auto dest!: Fun_canonical) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1855 | from T_App have "val t\<^isub>2 \<or> (\<exists>t'. t\<^isub>2 \<longmapsto> t')" by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1856 | thus ?case | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1857 | proof | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1858 | assume "val t\<^isub>2" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1859 | with t\<^isub>1 have "t\<^isub>1 \<cdot> t\<^isub>2 \<longmapsto> t3[x \<mapsto> t\<^isub>2]" by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1860 | thus ?case by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1861 | next | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1862 | assume "\<exists>t'. t\<^isub>2 \<longmapsto> t'" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1863 | then obtain t' where "t\<^isub>2 \<longmapsto> t'" by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1864 | with t\<^isub>1_val have "t\<^isub>1 \<cdot> t\<^isub>2 \<longmapsto> t\<^isub>1 \<cdot> t'" by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1865 | thus ?case by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1866 | qed | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1867 | next | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1868 | assume "\<exists>t'. t\<^isub>1 \<longmapsto> t'" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1869 | then obtain t' where "t\<^isub>1 \<longmapsto> t'" by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1870 | hence "t\<^isub>1 \<cdot> t\<^isub>2 \<longmapsto> t' \<cdot> t\<^isub>2" by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1871 | thus ?case by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1872 | qed | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1873 | next | 
| 35438 | 1874 | case (T_TApp X t\<^isub>1 T\<^isub>2 T\<^isub>1\<^isub>1 T\<^isub>1\<^isub>2) | 
| 30091 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1875 | hence "val t\<^isub>1 \<or> (\<exists>t'. t\<^isub>1 \<longmapsto> t')" by simp | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1876 | thus ?case | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1877 | proof | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1878 | assume "val t\<^isub>1" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1879 | with T_TApp obtain x t S where "t\<^isub>1 = (\<lambda>x<:S. t)" | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1880 | by (auto dest!: TyAll_canonical) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1881 | hence "t\<^isub>1 \<cdot>\<^sub>\<tau> T\<^isub>2 \<longmapsto> t[x \<mapsto>\<^sub>\<tau> T\<^isub>2]" by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1882 | thus ?case by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1883 | next | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1884 | assume "\<exists>t'. t\<^isub>1 \<longmapsto> t'" thus ?case by auto | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1885 | qed | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1886 | qed (auto) | 
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1887 | |
| 
2fb0b721e9c2
Added typing and evaluation relations, together with proofs of preservation
 berghofe parents: 
29097diff
changeset | 1888 | end |