| author | haftmann | 
| Fri, 07 May 2010 09:59:24 +0200 | |
| changeset 36725 | 34c36a5cb808 | 
| parent 32960 | 69916a850301 | 
| child 41229 | d797baa3d57c | 
| permissions | -rw-r--r-- | 
| 13163 | 1 | (* Title: ZF/func.thy | 
| 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 3 | Copyright 1991 University of Cambridge | |
| 4 | *) | |
| 5 | ||
| 13355 | 6 | header{*Functions, Function Spaces, Lambda-Abstraction*}
 | 
| 7 | ||
| 16417 | 8 | theory func imports equalities Sum begin | 
| 13163 | 9 | |
| 13355 | 10 | subsection{*The Pi Operator: Dependent Function Space*}
 | 
| 11 | ||
| 13248 | 12 | lemma subset_Sigma_imp_relation: "r <= Sigma(A,B) ==> relation(r)" | 
| 13 | by (simp add: relation_def, blast) | |
| 14 | ||
| 13221 | 15 | lemma relation_converse_converse [simp]: | 
| 16 | "relation(r) ==> converse(converse(r)) = r" | |
| 17 | by (simp add: relation_def, blast) | |
| 18 | ||
| 19 | lemma relation_restrict [simp]: "relation(restrict(r,A))" | |
| 20 | by (simp add: restrict_def relation_def, blast) | |
| 21 | ||
| 13163 | 22 | lemma Pi_iff: | 
| 23 | "f: Pi(A,B) <-> function(f) & f<=Sigma(A,B) & A<=domain(f)" | |
| 24 | by (unfold Pi_def, blast) | |
| 25 | ||
| 26 | (*For upward compatibility with the former definition*) | |
| 27 | lemma Pi_iff_old: | |
| 28 | "f: Pi(A,B) <-> f<=Sigma(A,B) & (ALL x:A. EX! y. <x,y>: f)" | |
| 29 | by (unfold Pi_def function_def, blast) | |
| 30 | ||
| 31 | lemma fun_is_function: "f: Pi(A,B) ==> function(f)" | |
| 32 | by (simp only: Pi_iff) | |
| 33 | ||
| 13219 | 34 | lemma function_imp_Pi: | 
| 35 | "[|function(f); relation(f)|] ==> f \<in> domain(f) -> range(f)" | |
| 36 | by (simp add: Pi_iff relation_def, blast) | |
| 37 | ||
| 13172 | 38 | lemma functionI: | 
| 39 | "[| !!x y y'. [| <x,y>:r; <x,y'>:r |] ==> y=y' |] ==> function(r)" | |
| 40 | by (simp add: function_def, blast) | |
| 41 | ||
| 13163 | 42 | (*Functions are relations*) | 
| 43 | lemma fun_is_rel: "f: Pi(A,B) ==> f <= Sigma(A,B)" | |
| 44 | by (unfold Pi_def, blast) | |
| 45 | ||
| 46 | lemma Pi_cong: | |
| 47 | "[| A=A'; !!x. x:A' ==> B(x)=B'(x) |] ==> Pi(A,B) = Pi(A',B')" | |
| 48 | by (simp add: Pi_def cong add: Sigma_cong) | |
| 49 | ||
| 50 | (*Sigma_cong, Pi_cong NOT given to Addcongs: they cause | |
| 51 | flex-flex pairs and the "Check your prover" error. Most | |
| 52 | Sigmas and Pis are abbreviated as * or -> *) | |
| 53 | ||
| 54 | (*Weakening one function type to another; see also Pi_type*) | |
| 55 | lemma fun_weaken_type: "[| f: A->B; B<=D |] ==> f: A->D" | |
| 56 | by (unfold Pi_def, best) | |
| 57 | ||
| 13355 | 58 | subsection{*Function Application*}
 | 
| 13163 | 59 | |
| 60 | lemma apply_equality2: "[| <a,b>: f; <a,c>: f; f: Pi(A,B) |] ==> b=c" | |
| 61 | by (unfold Pi_def function_def, blast) | |
| 62 | ||
| 63 | lemma function_apply_equality: "[| <a,b>: f; function(f) |] ==> f`a = b" | |
| 64 | by (unfold apply_def function_def, blast) | |
| 65 | ||
| 66 | lemma apply_equality: "[| <a,b>: f; f: Pi(A,B) |] ==> f`a = b" | |
| 67 | apply (unfold Pi_def) | |
| 68 | apply (blast intro: function_apply_equality) | |
| 69 | done | |
| 70 | ||
| 71 | (*Applying a function outside its domain yields 0*) | |
| 72 | lemma apply_0: "a ~: domain(f) ==> f`a = 0" | |
| 13176 
312bd350579b
conversion of Perm to Isar.  Strengthening of comp_fun_apply
 paulson parents: 
13175diff
changeset | 73 | by (unfold apply_def, blast) | 
| 13163 | 74 | |
| 75 | lemma Pi_memberD: "[| f: Pi(A,B); c: f |] ==> EX x:A. c = <x,f`x>" | |
| 76 | apply (frule fun_is_rel) | |
| 77 | apply (blast dest: apply_equality) | |
| 78 | done | |
| 79 | ||
| 80 | lemma function_apply_Pair: "[| function(f); a : domain(f)|] ==> <a,f`a>: f" | |
| 13176 
312bd350579b
conversion of Perm to Isar.  Strengthening of comp_fun_apply
 paulson parents: 
13175diff
changeset | 81 | apply (simp add: function_def, clarify) | 
| 
312bd350579b
conversion of Perm to Isar.  Strengthening of comp_fun_apply
 paulson parents: 
13175diff
changeset | 82 | apply (subgoal_tac "f`a = y", blast) | 
| 
312bd350579b
conversion of Perm to Isar.  Strengthening of comp_fun_apply
 paulson parents: 
13175diff
changeset | 83 | apply (simp add: apply_def, blast) | 
| 13163 | 84 | done | 
| 85 | ||
| 86 | lemma apply_Pair: "[| f: Pi(A,B); a:A |] ==> <a,f`a>: f" | |
| 87 | apply (simp add: Pi_iff) | |
| 88 | apply (blast intro: function_apply_Pair) | |
| 89 | done | |
| 90 | ||
| 27150 | 91 | (*Conclusion is flexible -- use rule_tac or else apply_funtype below!*) | 
| 13163 | 92 | lemma apply_type [TC]: "[| f: Pi(A,B); a:A |] ==> f`a : B(a)" | 
| 93 | by (blast intro: apply_Pair dest: fun_is_rel) | |
| 94 | ||
| 95 | (*This version is acceptable to the simplifier*) | |
| 96 | lemma apply_funtype: "[| f: A->B; a:A |] ==> f`a : B" | |
| 97 | by (blast dest: apply_type) | |
| 98 | ||
| 99 | lemma apply_iff: "f: Pi(A,B) ==> <a,b>: f <-> a:A & f`a = b" | |
| 100 | apply (frule fun_is_rel) | |
| 101 | apply (blast intro!: apply_Pair apply_equality) | |
| 102 | done | |
| 103 | ||
| 104 | (*Refining one Pi type to another*) | |
| 105 | lemma Pi_type: "[| f: Pi(A,C); !!x. x:A ==> f`x : B(x) |] ==> f : Pi(A,B)" | |
| 106 | apply (simp only: Pi_iff) | |
| 107 | apply (blast dest: function_apply_equality) | |
| 108 | done | |
| 109 | ||
| 110 | (*Such functions arise in non-standard datatypes, ZF/ex/Ntree for instance*) | |
| 111 | lemma Pi_Collect_iff: | |
| 112 |      "(f : Pi(A, %x. {y:B(x). P(x,y)}))
 | |
| 113 | <-> f : Pi(A,B) & (ALL x: A. P(x, f`x))" | |
| 114 | by (blast intro: Pi_type dest: apply_type) | |
| 115 | ||
| 116 | lemma Pi_weaken_type: | |
| 117 | "[| f : Pi(A,B); !!x. x:A ==> B(x)<=C(x) |] ==> f : Pi(A,C)" | |
| 118 | by (blast intro: Pi_type dest: apply_type) | |
| 119 | ||
| 120 | ||
| 121 | (** Elimination of membership in a function **) | |
| 122 | ||
| 123 | lemma domain_type: "[| <a,b> : f; f: Pi(A,B) |] ==> a : A" | |
| 124 | by (blast dest: fun_is_rel) | |
| 125 | ||
| 126 | lemma range_type: "[| <a,b> : f; f: Pi(A,B) |] ==> b : B(a)" | |
| 127 | by (blast dest: fun_is_rel) | |
| 128 | ||
| 129 | lemma Pair_mem_PiD: "[| <a,b>: f; f: Pi(A,B) |] ==> a:A & b:B(a) & f`a = b" | |
| 130 | by (blast intro: domain_type range_type apply_equality) | |
| 131 | ||
| 13355 | 132 | subsection{*Lambda Abstraction*}
 | 
| 13163 | 133 | |
| 134 | lemma lamI: "a:A ==> <a,b(a)> : (lam x:A. b(x))" | |
| 135 | apply (unfold lam_def) | |
| 136 | apply (erule RepFunI) | |
| 137 | done | |
| 138 | ||
| 139 | lemma lamE: | |
| 140 | "[| p: (lam x:A. b(x)); !!x.[| x:A; p=<x,b(x)> |] ==> P | |
| 141 | |] ==> P" | |
| 142 | by (simp add: lam_def, blast) | |
| 143 | ||
| 144 | lemma lamD: "[| <a,c>: (lam x:A. b(x)) |] ==> c = b(a)" | |
| 145 | by (simp add: lam_def) | |
| 146 | ||
| 147 | lemma lam_type [TC]: | |
| 148 | "[| !!x. x:A ==> b(x): B(x) |] ==> (lam x:A. b(x)) : Pi(A,B)" | |
| 149 | by (simp add: lam_def Pi_def function_def, blast) | |
| 150 | ||
| 151 | lemma lam_funtype: "(lam x:A. b(x)) : A -> {b(x). x:A}"
 | |
| 13176 
312bd350579b
conversion of Perm to Isar.  Strengthening of comp_fun_apply
 paulson parents: 
13175diff
changeset | 152 | by (blast intro: lam_type) | 
| 13163 | 153 | |
| 13172 | 154 | lemma function_lam: "function (lam x:A. b(x))" | 
| 155 | by (simp add: function_def lam_def) | |
| 156 | ||
| 157 | lemma relation_lam: "relation (lam x:A. b(x))" | |
| 158 | by (simp add: relation_def lam_def) | |
| 159 | ||
| 13175 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 paulson parents: 
13174diff
changeset | 160 | lemma beta_if [simp]: "(lam x:A. b(x)) ` a = (if a : A then b(a) else 0)" | 
| 13176 
312bd350579b
conversion of Perm to Isar.  Strengthening of comp_fun_apply
 paulson parents: 
13175diff
changeset | 161 | by (simp add: apply_def lam_def, blast) | 
| 13175 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 paulson parents: 
13174diff
changeset | 162 | |
| 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 paulson parents: 
13174diff
changeset | 163 | lemma beta: "a : A ==> (lam x:A. b(x)) ` a = b(a)" | 
| 13176 
312bd350579b
conversion of Perm to Isar.  Strengthening of comp_fun_apply
 paulson parents: 
13175diff
changeset | 164 | by (simp add: apply_def lam_def, blast) | 
| 13163 | 165 | |
| 166 | lemma lam_empty [simp]: "(lam x:0. b(x)) = 0" | |
| 167 | by (simp add: lam_def) | |
| 168 | ||
| 169 | lemma domain_lam [simp]: "domain(Lambda(A,b)) = A" | |
| 170 | by (simp add: lam_def, blast) | |
| 171 | ||
| 172 | (*congruence rule for lambda abstraction*) | |
| 173 | lemma lam_cong [cong]: | |
| 174 | "[| A=A'; !!x. x:A' ==> b(x)=b'(x) |] ==> Lambda(A,b) = Lambda(A',b')" | |
| 175 | by (simp only: lam_def cong add: RepFun_cong) | |
| 176 | ||
| 177 | lemma lam_theI: | |
| 178 | "(!!x. x:A ==> EX! y. Q(x,y)) ==> EX f. ALL x:A. Q(x, f`x)" | |
| 13175 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 paulson parents: 
13174diff
changeset | 179 | apply (rule_tac x = "lam x: A. THE y. Q (x,y)" in exI) | 
| 13176 
312bd350579b
conversion of Perm to Isar.  Strengthening of comp_fun_apply
 paulson parents: 
13175diff
changeset | 180 | apply simp | 
| 13163 | 181 | apply (blast intro: theI) | 
| 182 | done | |
| 183 | ||
| 184 | lemma lam_eqE: "[| (lam x:A. f(x)) = (lam x:A. g(x)); a:A |] ==> f(a)=g(a)" | |
| 185 | by (fast intro!: lamI elim: equalityE lamE) | |
| 186 | ||
| 187 | ||
| 188 | (*Empty function spaces*) | |
| 189 | lemma Pi_empty1 [simp]: "Pi(0,A) = {0}"
 | |
| 190 | by (unfold Pi_def function_def, blast) | |
| 191 | ||
| 192 | (*The singleton function*) | |
| 193 | lemma singleton_fun [simp]: "{<a,b>} : {a} -> {b}"
 | |
| 194 | by (unfold Pi_def function_def, blast) | |
| 195 | ||
| 196 | lemma Pi_empty2 [simp]: "(A->0) = (if A=0 then {0} else 0)"
 | |
| 197 | by (unfold Pi_def function_def, force) | |
| 198 | ||
| 199 | lemma fun_space_empty_iff [iff]: "(A->X)=0 \<longleftrightarrow> X=0 & (A \<noteq> 0)" | |
| 200 | apply auto | |
| 201 | apply (fast intro!: equals0I intro: lam_type) | |
| 202 | done | |
| 203 | ||
| 204 | ||
| 13355 | 205 | subsection{*Extensionality*}
 | 
| 13163 | 206 | |
| 207 | (*Semi-extensionality!*) | |
| 208 | ||
| 209 | lemma fun_subset: | |
| 210 | "[| f : Pi(A,B); g: Pi(C,D); A<=C; | |
| 211 | !!x. x:A ==> f`x = g`x |] ==> f<=g" | |
| 212 | by (force dest: Pi_memberD intro: apply_Pair) | |
| 213 | ||
| 214 | lemma fun_extension: | |
| 215 | "[| f : Pi(A,B); g: Pi(A,D); | |
| 216 | !!x. x:A ==> f`x = g`x |] ==> f=g" | |
| 217 | by (blast del: subsetI intro: subset_refl sym fun_subset) | |
| 218 | ||
| 219 | lemma eta [simp]: "f : Pi(A,B) ==> (lam x:A. f`x) = f" | |
| 220 | apply (rule fun_extension) | |
| 221 | apply (auto simp add: lam_type apply_type beta) | |
| 222 | done | |
| 223 | ||
| 224 | lemma fun_extension_iff: | |
| 225 | "[| f:Pi(A,B); g:Pi(A,C) |] ==> (ALL a:A. f`a = g`a) <-> f=g" | |
| 226 | by (blast intro: fun_extension) | |
| 227 | ||
| 228 | (*thm by Mark Staples, proof by lcp*) | |
| 229 | lemma fun_subset_eq: "[| f:Pi(A,B); g:Pi(A,C) |] ==> f <= g <-> (f = g)" | |
| 230 | by (blast dest: apply_Pair | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
27702diff
changeset | 231 | intro: fun_extension apply_equality [symmetric]) | 
| 13163 | 232 | |
| 233 | ||
| 234 | (*Every element of Pi(A,B) may be expressed as a lambda abstraction!*) | |
| 235 | lemma Pi_lamE: | |
| 236 | assumes major: "f: Pi(A,B)" | |
| 237 | and minor: "!!b. [| ALL x:A. b(x):B(x); f = (lam x:A. b(x)) |] ==> P" | |
| 238 | shows "P" | |
| 239 | apply (rule minor) | |
| 240 | apply (rule_tac [2] eta [symmetric]) | |
| 241 | apply (blast intro: major apply_type)+ | |
| 242 | done | |
| 243 | ||
| 244 | ||
| 13355 | 245 | subsection{*Images of Functions*}
 | 
| 13163 | 246 | |
| 247 | lemma image_lam: "C <= A ==> (lam x:A. b(x)) `` C = {b(x). x:C}"
 | |
| 248 | by (unfold lam_def, blast) | |
| 249 | ||
| 13179 | 250 | lemma Repfun_function_if: | 
| 251 | "function(f) | |
| 252 |       ==> {f`x. x:C} = (if C <= domain(f) then f``C else cons(0,f``C))";
 | |
| 253 | apply simp | |
| 254 | apply (intro conjI impI) | |
| 255 | apply (blast dest: function_apply_equality intro: function_apply_Pair) | |
| 256 | apply (rule equalityI) | |
| 257 | apply (blast intro!: function_apply_Pair apply_0) | |
| 258 | apply (blast dest: function_apply_equality intro: apply_0 [symmetric]) | |
| 259 | done | |
| 260 | ||
| 261 | (*For this lemma and the next, the right-hand side could equivalently | |
| 13615 
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
 paulson parents: 
13357diff
changeset | 262 |   be written \<Union>x\<in>C. {f`x} *)
 | 
| 13174 | 263 | lemma image_function: | 
| 264 |      "[| function(f);  C <= domain(f) |] ==> f``C = {f`x. x:C}";
 | |
| 13179 | 265 | by (simp add: Repfun_function_if) | 
| 13174 | 266 | |
| 13163 | 267 | lemma image_fun: "[| f : Pi(A,B);  C <= A |] ==> f``C = {f`x. x:C}"
 | 
| 13174 | 268 | apply (simp add: Pi_iff) | 
| 269 | apply (blast intro: image_function) | |
| 13163 | 270 | done | 
| 271 | ||
| 14883 | 272 | lemma image_eq_UN: | 
| 273 |   assumes f: "f \<in> Pi(A,B)" "C \<subseteq> A" shows "f``C = (\<Union>x\<in>C. {f ` x})"
 | |
| 274 | by (auto simp add: image_fun [OF f]) | |
| 275 | ||
| 13163 | 276 | lemma Pi_image_cons: | 
| 277 | "[| f: Pi(A,B); x: A |] ==> f `` cons(x,y) = cons(f`x, f``y)" | |
| 278 | by (blast dest: apply_equality apply_Pair) | |
| 279 | ||
| 124 | 280 | |
| 13355 | 281 | subsection{*Properties of @{term "restrict(f,A)"}*}
 | 
| 13163 | 282 | |
| 13179 | 283 | lemma restrict_subset: "restrict(f,A) <= f" | 
| 284 | by (unfold restrict_def, blast) | |
| 13163 | 285 | |
| 286 | lemma function_restrictI: | |
| 287 | "function(f) ==> function(restrict(f,A))" | |
| 288 | by (unfold restrict_def function_def, blast) | |
| 289 | ||
| 290 | lemma restrict_type2: "[| f: Pi(C,B); A<=C |] ==> restrict(f,A) : Pi(A,B)" | |
| 291 | by (simp add: Pi_iff function_def restrict_def, blast) | |
| 292 | ||
| 13179 | 293 | lemma restrict: "restrict(f,A) ` a = (if a : A then f`a else 0)" | 
| 13176 
312bd350579b
conversion of Perm to Isar.  Strengthening of comp_fun_apply
 paulson parents: 
13175diff
changeset | 294 | by (simp add: apply_def restrict_def, blast) | 
| 13163 | 295 | |
| 296 | lemma restrict_empty [simp]: "restrict(f,0) = 0" | |
| 13179 | 297 | by (unfold restrict_def, simp) | 
| 13163 | 298 | |
| 13172 | 299 | lemma restrict_iff: "z \<in> restrict(r,A) \<longleftrightarrow> z \<in> r & (\<exists>x\<in>A. \<exists>y. z = \<langle>x, y\<rangle>)" | 
| 300 | by (simp add: restrict_def) | |
| 301 | ||
| 13163 | 302 | lemma restrict_restrict [simp]: | 
| 303 | "restrict(restrict(r,A),B) = restrict(r, A Int B)" | |
| 304 | by (unfold restrict_def, blast) | |
| 305 | ||
| 306 | lemma domain_restrict [simp]: "domain(restrict(f,C)) = domain(f) Int C" | |
| 307 | apply (unfold restrict_def) | |
| 308 | apply (auto simp add: domain_def) | |
| 309 | done | |
| 310 | ||
| 13248 | 311 | lemma restrict_idem: "f <= Sigma(A,B) ==> restrict(f,A) = f" | 
| 13163 | 312 | by (simp add: restrict_def, blast) | 
| 313 | ||
| 13248 | 314 | |
| 315 | (*converse probably holds too*) | |
| 316 | lemma domain_restrict_idem: | |
| 317 | "[| domain(r) <= A; relation(r) |] ==> restrict(r,A) = r" | |
| 318 | by (simp add: restrict_def relation_def, blast) | |
| 319 | ||
| 320 | lemma domain_restrict_lam [simp]: "domain(restrict(Lambda(A,f),C)) = A Int C" | |
| 321 | apply (unfold restrict_def lam_def) | |
| 322 | apply (rule equalityI) | |
| 323 | apply (auto simp add: domain_iff) | |
| 324 | done | |
| 325 | ||
| 13163 | 326 | lemma restrict_if [simp]: "restrict(f,A) ` a = (if a : A then f`a else 0)" | 
| 327 | by (simp add: restrict apply_0) | |
| 328 | ||
| 329 | lemma restrict_lam_eq: | |
| 330 | "A<=C ==> restrict(lam x:C. b(x), A) = (lam x:A. b(x))" | |
| 331 | by (unfold restrict_def lam_def, auto) | |
| 332 | ||
| 333 | lemma fun_cons_restrict_eq: | |
| 334 | "f : cons(a, b) -> B ==> f = cons(<a, f ` a>, restrict(f, b))" | |
| 335 | apply (rule equalityI) | |
| 13248 | 336 | prefer 2 apply (blast intro: apply_Pair restrict_subset [THEN subsetD]) | 
| 13163 | 337 | apply (auto dest!: Pi_memberD simp add: restrict_def lam_def) | 
| 338 | done | |
| 339 | ||
| 340 | ||
| 13355 | 341 | subsection{*Unions of Functions*}
 | 
| 13163 | 342 | |
| 343 | (** The Union of a set of COMPATIBLE functions is a function **) | |
| 344 | ||
| 345 | lemma function_Union: | |
| 346 | "[| ALL x:S. function(x); | |
| 347 | ALL x:S. ALL y:S. x<=y | y<=x |] | |
| 348 | ==> function(Union(S))" | |
| 349 | by (unfold function_def, blast) | |
| 350 | ||
| 351 | lemma fun_Union: | |
| 352 | "[| ALL f:S. EX C D. f:C->D; | |
| 353 | ALL f:S. ALL y:S. f<=y | y<=f |] ==> | |
| 354 | Union(S) : domain(Union(S)) -> range(Union(S))" | |
| 355 | apply (unfold Pi_def) | |
| 356 | apply (blast intro!: rel_Union function_Union) | |
| 357 | done | |
| 358 | ||
| 13174 | 359 | lemma gen_relation_Union [rule_format]: | 
| 360 | "\<forall>f\<in>F. relation(f) \<Longrightarrow> relation(Union(F))" | |
| 361 | by (simp add: relation_def) | |
| 362 | ||
| 13163 | 363 | |
| 364 | (** The Union of 2 disjoint functions is a function **) | |
| 365 | ||
| 366 | lemmas Un_rls = Un_subset_iff SUM_Un_distrib1 prod_Un_distrib2 | |
| 367 | subset_trans [OF _ Un_upper1] | |
| 368 | subset_trans [OF _ Un_upper2] | |
| 369 | ||
| 370 | lemma fun_disjoint_Un: | |
| 371 | "[| f: A->B; g: C->D; A Int C = 0 |] | |
| 372 | ==> (f Un g) : (A Un C) -> (B Un D)" | |
| 373 | (*Prove the product and domain subgoals using distributive laws*) | |
| 374 | apply (simp add: Pi_iff extension Un_rls) | |
| 375 | apply (unfold function_def, blast) | |
| 376 | done | |
| 377 | ||
| 13179 | 378 | lemma fun_disjoint_apply1: "a \<notin> domain(g) ==> (f Un g)`a = f`a" | 
| 379 | by (simp add: apply_def, blast) | |
| 13163 | 380 | |
| 13179 | 381 | lemma fun_disjoint_apply2: "c \<notin> domain(f) ==> (f Un g)`c = g`c" | 
| 382 | by (simp add: apply_def, blast) | |
| 13163 | 383 | |
| 13355 | 384 | subsection{*Domain and Range of a Function or Relation*}
 | 
| 13163 | 385 | |
| 386 | lemma domain_of_fun: "f : Pi(A,B) ==> domain(f)=A" | |
| 387 | by (unfold Pi_def, blast) | |
| 388 | ||
| 389 | lemma apply_rangeI: "[| f : Pi(A,B); a: A |] ==> f`a : range(f)" | |
| 390 | by (erule apply_Pair [THEN rangeI], assumption) | |
| 391 | ||
| 392 | lemma range_of_fun: "f : Pi(A,B) ==> f : A->range(f)" | |
| 393 | by (blast intro: Pi_type apply_rangeI) | |
| 394 | ||
| 13355 | 395 | subsection{*Extensions of Functions*}
 | 
| 13163 | 396 | |
| 397 | lemma fun_extend: | |
| 398 | "[| f: A->B; c~:A |] ==> cons(<c,b>,f) : cons(c,A) -> cons(b,B)" | |
| 399 | apply (frule singleton_fun [THEN fun_disjoint_Un], blast) | |
| 400 | apply (simp add: cons_eq) | |
| 401 | done | |
| 402 | ||
| 403 | lemma fun_extend3: | |
| 404 | "[| f: A->B; c~:A; b: B |] ==> cons(<c,b>,f) : cons(c,A) -> B" | |
| 405 | by (blast intro: fun_extend [THEN fun_weaken_type]) | |
| 406 | ||
| 13176 
312bd350579b
conversion of Perm to Isar.  Strengthening of comp_fun_apply
 paulson parents: 
13175diff
changeset | 407 | lemma extend_apply: | 
| 
312bd350579b
conversion of Perm to Isar.  Strengthening of comp_fun_apply
 paulson parents: 
13175diff
changeset | 408 | "c ~: domain(f) ==> cons(<c,b>,f)`a = (if a=c then b else f`a)" | 
| 
312bd350579b
conversion of Perm to Isar.  Strengthening of comp_fun_apply
 paulson parents: 
13175diff
changeset | 409 | by (auto simp add: apply_def) | 
| 13163 | 410 | |
| 13176 
312bd350579b
conversion of Perm to Isar.  Strengthening of comp_fun_apply
 paulson parents: 
13175diff
changeset | 411 | lemma fun_extend_apply [simp]: | 
| 
312bd350579b
conversion of Perm to Isar.  Strengthening of comp_fun_apply
 paulson parents: 
13175diff
changeset | 412 | "[| f: A->B; c~:A |] ==> cons(<c,b>,f)`a = (if a=c then b else f`a)" | 
| 
312bd350579b
conversion of Perm to Isar.  Strengthening of comp_fun_apply
 paulson parents: 
13175diff
changeset | 413 | apply (rule extend_apply) | 
| 
312bd350579b
conversion of Perm to Isar.  Strengthening of comp_fun_apply
 paulson parents: 
13175diff
changeset | 414 | apply (simp add: Pi_def, blast) | 
| 13163 | 415 | done | 
| 416 | ||
| 417 | lemmas singleton_apply = apply_equality [OF singletonI singleton_fun, simp] | |
| 418 | ||
| 419 | (*For Finite.ML. Inclusion of right into left is easy*) | |
| 420 | lemma cons_fun_eq: | |
| 13269 | 421 |      "c ~: A ==> cons(c,A) -> B = (\<Union>f \<in> A->B. \<Union>b\<in>B. {cons(<c,b>, f)})"
 | 
| 13163 | 422 | apply (rule equalityI) | 
| 423 | apply (safe elim!: fun_extend3) | |
| 424 | (*Inclusion of left into right*) | |
| 425 | apply (subgoal_tac "restrict (x, A) : A -> B") | |
| 426 | prefer 2 apply (blast intro: restrict_type2) | |
| 427 | apply (rule UN_I, assumption) | |
| 428 | apply (rule apply_funtype [THEN UN_I]) | |
| 429 | apply assumption | |
| 430 | apply (rule consI1) | |
| 431 | apply (simp (no_asm)) | |
| 432 | apply (rule fun_extension) | |
| 433 | apply assumption | |
| 434 | apply (blast intro: fun_extend) | |
| 13176 
312bd350579b
conversion of Perm to Isar.  Strengthening of comp_fun_apply
 paulson parents: 
13175diff
changeset | 435 | apply (erule consE, simp_all) | 
| 13163 | 436 | done | 
| 437 | ||
| 13269 | 438 | lemma succ_fun_eq: "succ(n) -> B = (\<Union>f \<in> n->B. \<Union>b\<in>B. {cons(<n,b>, f)})"
 | 
| 439 | by (simp add: succ_def mem_not_refl cons_fun_eq) | |
| 440 | ||
| 13355 | 441 | |
| 442 | subsection{*Function Updates*}
 | |
| 443 | ||
| 24893 | 444 | definition | 
| 445 | update :: "[i,i,i] => i" where | |
| 13355 | 446 | "update(f,a,b) == lam x: cons(a, domain(f)). if(x=a, b, f`x)" | 
| 447 | ||
| 448 | nonterminals | |
| 449 | updbinds updbind | |
| 450 | ||
| 451 | syntax | |
| 452 | ||
| 453 | (* Let expressions *) | |
| 454 | ||
| 455 |   "_updbind"    :: "[i, i] => updbind"               ("(2_ :=/ _)")
 | |
| 456 |   ""            :: "updbind => updbinds"             ("_")
 | |
| 457 |   "_updbinds"   :: "[updbind, updbinds] => updbinds" ("_,/ _")
 | |
| 458 |   "_Update"     :: "[i, updbinds] => i"              ("_/'((_)')" [900,0] 900)
 | |
| 459 | ||
| 460 | translations | |
| 461 | "_Update (f, _updbinds(b,bs))" == "_Update (_Update(f,b), bs)" | |
| 24893 | 462 | "f(x:=y)" == "CONST update(f,x,y)" | 
| 13355 | 463 | |
| 464 | ||
| 465 | lemma update_apply [simp]: "f(x:=y) ` z = (if z=x then y else f`z)" | |
| 466 | apply (simp add: update_def) | |
| 14153 | 467 | apply (case_tac "z \<in> domain(f)") | 
| 13355 | 468 | apply (simp_all add: apply_0) | 
| 469 | done | |
| 470 | ||
| 471 | lemma update_idem: "[| f`x = y; f: Pi(A,B); x: A |] ==> f(x:=y) = f" | |
| 472 | apply (unfold update_def) | |
| 473 | apply (simp add: domain_of_fun cons_absorb) | |
| 474 | apply (rule fun_extension) | |
| 475 | apply (best intro: apply_type if_type lam_type, assumption, simp) | |
| 476 | done | |
| 477 | ||
| 478 | ||
| 479 | (* [| f: Pi(A, B); x:A |] ==> f(x := f`x) = f *) | |
| 480 | declare refl [THEN update_idem, simp] | |
| 481 | ||
| 482 | lemma domain_update [simp]: "domain(f(x:=y)) = cons(x, domain(f))" | |
| 483 | by (unfold update_def, simp) | |
| 484 | ||
| 14046 | 485 | lemma update_type: "[| f:Pi(A,B); x : A; y: B(x) |] ==> f(x:=y) : Pi(A, B)" | 
| 13355 | 486 | apply (unfold update_def) | 
| 487 | apply (simp add: domain_of_fun cons_absorb apply_funtype lam_type) | |
| 488 | done | |
| 489 | ||
| 490 | ||
| 13357 | 491 | subsection{*Monotonicity Theorems*}
 | 
| 492 | ||
| 493 | subsubsection{*Replacement in its Various Forms*}
 | |
| 494 | ||
| 495 | (*Not easy to express monotonicity in P, since any "bigger" predicate | |
| 496 | would have to be single-valued*) | |
| 497 | lemma Replace_mono: "A<=B ==> Replace(A,P) <= Replace(B,P)" | |
| 498 | by (blast elim!: ReplaceE) | |
| 499 | ||
| 500 | lemma RepFun_mono: "A<=B ==> {f(x). x:A} <= {f(x). x:B}"
 | |
| 501 | by blast | |
| 502 | ||
| 503 | lemma Pow_mono: "A<=B ==> Pow(A) <= Pow(B)" | |
| 504 | by blast | |
| 505 | ||
| 506 | lemma Union_mono: "A<=B ==> Union(A) <= Union(B)" | |
| 507 | by blast | |
| 508 | ||
| 509 | lemma UN_mono: | |
| 13615 
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
 paulson parents: 
13357diff
changeset | 510 | "[| A<=C; !!x. x:A ==> B(x)<=D(x) |] ==> (\<Union>x\<in>A. B(x)) <= (\<Union>x\<in>C. D(x))" | 
| 13357 | 511 | by blast | 
| 512 | ||
| 513 | (*Intersection is ANTI-monotonic. There are TWO premises! *) | |
| 14095 
a1ba833d6b61
Changed many Intersection rules from i:I to I~=0 to avoid introducing a new
 paulson parents: 
14046diff
changeset | 514 | lemma Inter_anti_mono: "[| A<=B; A\<noteq>0 |] ==> Inter(B) <= Inter(A)" | 
| 13357 | 515 | by blast | 
| 516 | ||
| 517 | lemma cons_mono: "C<=D ==> cons(a,C) <= cons(a,D)" | |
| 518 | by blast | |
| 519 | ||
| 520 | lemma Un_mono: "[| A<=C; B<=D |] ==> A Un B <= C Un D" | |
| 521 | by blast | |
| 522 | ||
| 523 | lemma Int_mono: "[| A<=C; B<=D |] ==> A Int B <= C Int D" | |
| 524 | by blast | |
| 525 | ||
| 526 | lemma Diff_mono: "[| A<=C; D<=B |] ==> A-B <= C-D" | |
| 527 | by blast | |
| 528 | ||
| 529 | subsubsection{*Standard Products, Sums and Function Spaces *}
 | |
| 530 | ||
| 531 | lemma Sigma_mono [rule_format]: | |
| 532 | "[| A<=C; !!x. x:A --> B(x) <= D(x) |] ==> Sigma(A,B) <= Sigma(C,D)" | |
| 533 | by blast | |
| 534 | ||
| 535 | lemma sum_mono: "[| A<=C; B<=D |] ==> A+B <= C+D" | |
| 536 | by (unfold sum_def, blast) | |
| 537 | ||
| 538 | (*Note that B->A and C->A are typically disjoint!*) | |
| 539 | lemma Pi_mono: "B<=C ==> A->B <= A->C" | |
| 540 | by (blast intro: lam_type elim: Pi_lamE) | |
| 541 | ||
| 542 | lemma lam_mono: "A<=B ==> Lambda(A,c) <= Lambda(B,c)" | |
| 543 | apply (unfold lam_def) | |
| 544 | apply (erule RepFun_mono) | |
| 545 | done | |
| 546 | ||
| 547 | subsubsection{*Converse, Domain, Range, Field*}
 | |
| 548 | ||
| 549 | lemma converse_mono: "r<=s ==> converse(r) <= converse(s)" | |
| 550 | by blast | |
| 551 | ||
| 552 | lemma domain_mono: "r<=s ==> domain(r)<=domain(s)" | |
| 553 | by blast | |
| 554 | ||
| 555 | lemmas domain_rel_subset = subset_trans [OF domain_mono domain_subset] | |
| 556 | ||
| 557 | lemma range_mono: "r<=s ==> range(r)<=range(s)" | |
| 558 | by blast | |
| 559 | ||
| 560 | lemmas range_rel_subset = subset_trans [OF range_mono range_subset] | |
| 561 | ||
| 562 | lemma field_mono: "r<=s ==> field(r)<=field(s)" | |
| 563 | by blast | |
| 564 | ||
| 565 | lemma field_rel_subset: "r <= A*A ==> field(r) <= A" | |
| 566 | by (erule field_mono [THEN subset_trans], blast) | |
| 567 | ||
| 568 | ||
| 569 | subsubsection{*Images*}
 | |
| 570 | ||
| 571 | lemma image_pair_mono: | |
| 572 | "[| !! x y. <x,y>:r ==> <x,y>:s; A<=B |] ==> r``A <= s``B" | |
| 573 | by blast | |
| 574 | ||
| 575 | lemma vimage_pair_mono: | |
| 576 | "[| !! x y. <x,y>:r ==> <x,y>:s; A<=B |] ==> r-``A <= s-``B" | |
| 577 | by blast | |
| 578 | ||
| 579 | lemma image_mono: "[| r<=s; A<=B |] ==> r``A <= s``B" | |
| 580 | by blast | |
| 581 | ||
| 582 | lemma vimage_mono: "[| r<=s; A<=B |] ==> r-``A <= s-``B" | |
| 583 | by blast | |
| 584 | ||
| 585 | lemma Collect_mono: | |
| 586 | "[| A<=B; !!x. x:A ==> P(x) --> Q(x) |] ==> Collect(A,P) <= Collect(B,Q)" | |
| 587 | by blast | |
| 588 | ||
| 589 | (*Used in intr_elim.ML and in individual datatype definitions*) | |
| 590 | lemmas basic_monos = subset_refl imp_refl disj_mono conj_mono ex_mono | |
| 591 | Collect_mono Part_mono in_mono | |
| 592 | ||
| 27702 | 593 | (* Useful with simp; contributed by Clemens Ballarin. *) | 
| 594 | ||
| 595 | lemma bex_image_simp: | |
| 596 | "[| f : Pi(X, Y); A \<subseteq> X |] ==> (EX x : f``A. P(x)) <-> (EX x:A. P(f`x))" | |
| 597 | apply safe | |
| 598 | apply rule | |
| 599 | prefer 2 apply assumption | |
| 600 | apply (simp add: apply_equality) | |
| 601 | apply (blast intro: apply_Pair) | |
| 602 | done | |
| 603 | ||
| 604 | lemma ball_image_simp: | |
| 605 | "[| f : Pi(X, Y); A \<subseteq> X |] ==> (ALL x : f``A. P(x)) <-> (ALL x:A. P(f`x))" | |
| 606 | apply safe | |
| 607 | apply (blast intro: apply_Pair) | |
| 608 | apply (drule bspec) apply assumption | |
| 609 | apply (simp add: apply_equality) | |
| 610 | done | |
| 611 | ||
| 13163 | 612 | end |