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