| author | paulson | 
| Wed, 09 Jun 2004 11:19:23 +0200 | |
| changeset 14890 | 51f28df21c8b | 
| parent 14565 | c6dc17aab88a | 
| child 15111 | c108189645f8 | 
| permissions | -rw-r--r-- | 
| 1475 | 1 | (* Title: HOL/Fun.thy | 
| 923 | 2 | ID: $Id$ | 
| 1475 | 3 | Author: Tobias Nipkow, Cambridge University Computer Laboratory | 
| 923 | 4 | Copyright 1994 University of Cambridge | 
| 5 | ||
| 2912 | 6 | Notions about functions. | 
| 923 | 7 | *) | 
| 8 | ||
| 13585 | 9 | theory Fun = Typedef: | 
| 2912 | 10 | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12258diff
changeset | 11 | instance set :: (type) order | 
| 13585 | 12 | by (intro_classes, | 
| 13 | (assumption | rule subset_refl subset_trans subset_antisym psubset_eq)+) | |
| 14 | ||
| 15 | constdefs | |
| 16 |   fun_upd :: "('a => 'b) => 'a => 'b => ('a => 'b)"
 | |
| 17 | "fun_upd f a b == % x. if x=a then b else f x" | |
| 6171 | 18 | |
| 9141 | 19 | nonterminals | 
| 20 | updbinds updbind | |
| 5305 | 21 | syntax | 
| 13585 | 22 |   "_updbind" :: "['a, 'a] => updbind"             ("(2_ :=/ _)")
 | 
| 23 |   ""         :: "updbind => updbinds"             ("_")
 | |
| 24 |   "_updbinds":: "[updbind, updbinds] => updbinds" ("_,/ _")
 | |
| 25 |   "_Update"  :: "['a, updbinds] => 'a"            ("_/'((_)')" [1000,0] 900)
 | |
| 5305 | 26 | |
| 27 | translations | |
| 28 | "_Update f (_updbinds b bs)" == "_Update (_Update f b) bs" | |
| 29 | "f(x:=y)" == "fun_upd f x y" | |
| 2912 | 30 | |
| 9340 | 31 | (* Hint: to define the sum of two functions (or maps), use sum_case. | 
| 32 | A nice infix syntax could be defined (in Datatype.thy or below) by | |
| 33 | consts | |
| 34 |   fun_sum :: "('a => 'c) => ('b => 'c) => (('a+'b) => 'c)" (infixr "'(+')"80)
 | |
| 35 | translations | |
| 13585 | 36 | "fun_sum" == sum_case | 
| 9340 | 37 | *) | 
| 12258 | 38 | |
| 6171 | 39 | constdefs | 
| 13910 | 40 |  overwrite :: "('a => 'b) => ('a => 'b) => 'a set => ('a => 'b)"
 | 
| 41 |               ("_/'(_|/_')"  [900,0,0]900)
 | |
| 42 | "f(g|A) == %a. if a : A then g a else f a" | |
| 6171 | 43 | |
| 13910 | 44 | id :: "'a => 'a" | 
| 45 | "id == %x. x" | |
| 46 | ||
| 47 | comp :: "['b => 'c, 'a => 'b, 'a] => 'c" (infixl "o" 55) | |
| 48 | "f o g == %x. f(g(x))" | |
| 11123 | 49 | |
| 13585 | 50 | text{*compatibility*}
 | 
| 51 | lemmas o_def = comp_def | |
| 2912 | 52 | |
| 12114 
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
 wenzelm parents: 
11609diff
changeset | 53 | syntax (xsymbols) | 
| 13585 | 54 | comp :: "['b => 'c, 'a => 'b, 'a] => 'c" (infixl "\<circ>" 55) | 
| 14565 | 55 | syntax (HTML output) | 
| 56 | comp :: "['b => 'c, 'a => 'b, 'a] => 'c" (infixl "\<circ>" 55) | |
| 13585 | 57 | |
| 9352 | 58 | |
| 13585 | 59 | constdefs | 
| 60 | inj_on :: "['a => 'b, 'a set] => bool" (*injective*) | |
| 61 | "inj_on f A == ! x:A. ! y:A. f(x)=f(y) --> x=y" | |
| 6171 | 62 | |
| 13585 | 63 | text{*A common special case: functions injective over the entire domain type.*}
 | 
| 64 | syntax inj   :: "('a => 'b) => bool"
 | |
| 6171 | 65 | translations | 
| 66 | "inj f" == "inj_on f UNIV" | |
| 5852 | 67 | |
| 7374 | 68 | constdefs | 
| 13585 | 69 |   surj :: "('a => 'b) => bool"                   (*surjective*)
 | 
| 7374 | 70 | "surj f == ! y. ? x. y=f(x)" | 
| 12258 | 71 | |
| 13585 | 72 |   bij :: "('a => 'b) => bool"                    (*bijective*)
 | 
| 7374 | 73 | "bij f == inj f & surj f" | 
| 12258 | 74 | |
| 7374 | 75 | |
| 13585 | 76 | |
| 77 | text{*As a simplification rule, it replaces all function equalities by
 | |
| 78 | first-order equalities.*} | |
| 79 | lemma expand_fun_eq: "(f = g) = (! x. f(x)=g(x))" | |
| 80 | apply (rule iffI) | |
| 81 | apply (simp (no_asm_simp)) | |
| 82 | apply (rule ext, simp (no_asm_simp)) | |
| 83 | done | |
| 84 | ||
| 85 | lemma apply_inverse: | |
| 86 | "[| f(x)=u; !!x. P(x) ==> g(f(x)) = x; P(x) |] ==> x=g(u)" | |
| 87 | by auto | |
| 88 | ||
| 89 | ||
| 90 | text{*The Identity Function: @{term id}*}
 | |
| 91 | lemma id_apply [simp]: "id x = x" | |
| 92 | by (simp add: id_def) | |
| 93 | ||
| 94 | ||
| 95 | subsection{*The Composition Operator: @{term "f \<circ> g"}*}
 | |
| 96 | ||
| 97 | lemma o_apply [simp]: "(f o g) x = f (g x)" | |
| 98 | by (simp add: comp_def) | |
| 99 | ||
| 100 | lemma o_assoc: "f o (g o h) = f o g o h" | |
| 101 | by (simp add: comp_def) | |
| 102 | ||
| 103 | lemma id_o [simp]: "id o g = g" | |
| 104 | by (simp add: comp_def) | |
| 105 | ||
| 106 | lemma o_id [simp]: "f o id = f" | |
| 107 | by (simp add: comp_def) | |
| 108 | ||
| 109 | lemma image_compose: "(f o g) ` r = f`(g`r)" | |
| 110 | by (simp add: comp_def, blast) | |
| 111 | ||
| 112 | lemma image_eq_UN: "f`A = (UN x:A. {f x})"
 | |
| 113 | by blast | |
| 114 | ||
| 115 | lemma UN_o: "UNION A (g o f) = UNION (f`A) g" | |
| 116 | by (unfold comp_def, blast) | |
| 117 | ||
| 118 | ||
| 119 | subsection{*The Injectivity Predicate, @{term inj}*}
 | |
| 120 | ||
| 121 | text{*NB: @{term inj} now just translates to @{term inj_on}*}
 | |
| 122 | ||
| 123 | ||
| 124 | text{*For Proofs in @{text "Tools/datatype_rep_proofs"}*}
 | |
| 125 | lemma datatype_injI: | |
| 126 | "(!! x. ALL y. f(x) = f(y) --> x=y) ==> inj(f)" | |
| 127 | by (simp add: inj_on_def) | |
| 128 | ||
| 13637 | 129 | theorem range_ex1_eq: "inj f \<Longrightarrow> b : range f = (EX! x. b = f x)" | 
| 130 | by (unfold inj_on_def, blast) | |
| 131 | ||
| 13585 | 132 | lemma injD: "[| inj(f); f(x) = f(y) |] ==> x=y" | 
| 133 | by (simp add: inj_on_def) | |
| 134 | ||
| 135 | (*Useful with the simplifier*) | |
| 136 | lemma inj_eq: "inj(f) ==> (f(x) = f(y)) = (x=y)" | |
| 137 | by (force simp add: inj_on_def) | |
| 138 | ||
| 139 | ||
| 140 | subsection{*The Predicate @{term inj_on}: Injectivity On A Restricted Domain*}
 | |
| 141 | ||
| 142 | lemma inj_onI: | |
| 143 | "(!! x y. [| x:A; y:A; f(x) = f(y) |] ==> x=y) ==> inj_on f A" | |
| 144 | by (simp add: inj_on_def) | |
| 145 | ||
| 146 | lemma inj_on_inverseI: "(!!x. x:A ==> g(f(x)) = x) ==> inj_on f A" | |
| 147 | by (auto dest: arg_cong [of concl: g] simp add: inj_on_def) | |
| 148 | ||
| 149 | lemma inj_onD: "[| inj_on f A; f(x)=f(y); x:A; y:A |] ==> x=y" | |
| 150 | by (unfold inj_on_def, blast) | |
| 151 | ||
| 152 | lemma inj_on_iff: "[| inj_on f A; x:A; y:A |] ==> (f(x)=f(y)) = (x=y)" | |
| 153 | by (blast dest!: inj_onD) | |
| 154 | ||
| 155 | lemma comp_inj_on: | |
| 156 | "[| inj_on f A; inj_on g (f`A) |] ==> inj_on (g o f) A" | |
| 157 | by (simp add: comp_def inj_on_def) | |
| 158 | ||
| 159 | lemma inj_on_contraD: "[| inj_on f A; ~x=y; x:A; y:A |] ==> ~ f(x)=f(y)" | |
| 160 | by (unfold inj_on_def, blast) | |
| 12258 | 161 | |
| 13585 | 162 | lemma inj_singleton: "inj (%s. {s})"
 | 
| 163 | by (simp add: inj_on_def) | |
| 164 | ||
| 165 | lemma subset_inj_on: "[| A<=B; inj_on f B |] ==> inj_on f A" | |
| 166 | by (unfold inj_on_def, blast) | |
| 167 | ||
| 168 | ||
| 169 | subsection{*The Predicate @{term surj}: Surjectivity*}
 | |
| 170 | ||
| 171 | lemma surjI: "(!! x. g(f x) = x) ==> surj g" | |
| 172 | apply (simp add: surj_def) | |
| 173 | apply (blast intro: sym) | |
| 174 | done | |
| 175 | ||
| 176 | lemma surj_range: "surj f ==> range f = UNIV" | |
| 177 | by (auto simp add: surj_def) | |
| 178 | ||
| 179 | lemma surjD: "surj f ==> EX x. y = f x" | |
| 180 | by (simp add: surj_def) | |
| 181 | ||
| 182 | lemma surjE: "surj f ==> (!!x. y = f x ==> C) ==> C" | |
| 183 | by (simp add: surj_def, blast) | |
| 184 | ||
| 185 | lemma comp_surj: "[| surj f; surj g |] ==> surj (g o f)" | |
| 186 | apply (simp add: comp_def surj_def, clarify) | |
| 187 | apply (drule_tac x = y in spec, clarify) | |
| 188 | apply (drule_tac x = x in spec, blast) | |
| 189 | done | |
| 190 | ||
| 191 | ||
| 192 | ||
| 193 | subsection{*The Predicate @{term bij}: Bijectivity*}
 | |
| 194 | ||
| 195 | lemma bijI: "[| inj f; surj f |] ==> bij f" | |
| 196 | by (simp add: bij_def) | |
| 197 | ||
| 198 | lemma bij_is_inj: "bij f ==> inj f" | |
| 199 | by (simp add: bij_def) | |
| 200 | ||
| 201 | lemma bij_is_surj: "bij f ==> surj f" | |
| 202 | by (simp add: bij_def) | |
| 203 | ||
| 204 | ||
| 205 | subsection{*Facts About the Identity Function*}
 | |
| 5852 | 206 | |
| 13585 | 207 | text{*We seem to need both the @{term id} forms and the @{term "\<lambda>x. x"}
 | 
| 208 | forms. The latter can arise by rewriting, while @{term id} may be used
 | |
| 209 | explicitly.*} | |
| 210 | ||
| 211 | lemma image_ident [simp]: "(%x. x) ` Y = Y" | |
| 212 | by blast | |
| 213 | ||
| 214 | lemma image_id [simp]: "id ` Y = Y" | |
| 215 | by (simp add: id_def) | |
| 216 | ||
| 217 | lemma vimage_ident [simp]: "(%x. x) -` Y = Y" | |
| 218 | by blast | |
| 219 | ||
| 220 | lemma vimage_id [simp]: "id -` A = A" | |
| 221 | by (simp add: id_def) | |
| 222 | ||
| 223 | lemma vimage_image_eq: "f -` (f ` A) = {y. EX x:A. f x = f y}"
 | |
| 224 | by (blast intro: sym) | |
| 225 | ||
| 226 | lemma image_vimage_subset: "f ` (f -` A) <= A" | |
| 227 | by blast | |
| 228 | ||
| 229 | lemma image_vimage_eq [simp]: "f ` (f -` A) = A Int range f" | |
| 230 | by blast | |
| 231 | ||
| 232 | lemma surj_image_vimage_eq: "surj f ==> f ` (f -` A) = A" | |
| 233 | by (simp add: surj_range) | |
| 234 | ||
| 235 | lemma inj_vimage_image_eq: "inj f ==> f -` (f ` A) = A" | |
| 236 | by (simp add: inj_on_def, blast) | |
| 237 | ||
| 238 | lemma vimage_subsetD: "surj f ==> f -` B <= A ==> B <= f ` A" | |
| 239 | apply (unfold surj_def) | |
| 240 | apply (blast intro: sym) | |
| 241 | done | |
| 242 | ||
| 243 | lemma vimage_subsetI: "inj f ==> B <= f ` A ==> f -` B <= A" | |
| 244 | by (unfold inj_on_def, blast) | |
| 245 | ||
| 246 | lemma vimage_subset_eq: "bij f ==> (f -` B <= A) = (B <= f ` A)" | |
| 247 | apply (unfold bij_def) | |
| 248 | apply (blast del: subsetI intro: vimage_subsetI vimage_subsetD) | |
| 249 | done | |
| 250 | ||
| 251 | lemma image_Int_subset: "f`(A Int B) <= f`A Int f`B" | |
| 252 | by blast | |
| 253 | ||
| 254 | lemma image_diff_subset: "f`A - f`B <= f`(A - B)" | |
| 255 | by blast | |
| 5852 | 256 | |
| 13585 | 257 | lemma inj_on_image_Int: | 
| 258 | "[| inj_on f C; A<=C; B<=C |] ==> f`(A Int B) = f`A Int f`B" | |
| 259 | apply (simp add: inj_on_def, blast) | |
| 260 | done | |
| 261 | ||
| 262 | lemma inj_on_image_set_diff: | |
| 263 | "[| inj_on f C; A<=C; B<=C |] ==> f`(A-B) = f`A - f`B" | |
| 264 | apply (simp add: inj_on_def, blast) | |
| 265 | done | |
| 266 | ||
| 267 | lemma image_Int: "inj f ==> f`(A Int B) = f`A Int f`B" | |
| 268 | by (simp add: inj_on_def, blast) | |
| 269 | ||
| 270 | lemma image_set_diff: "inj f ==> f`(A-B) = f`A - f`B" | |
| 271 | by (simp add: inj_on_def, blast) | |
| 272 | ||
| 273 | lemma inj_image_mem_iff: "inj f ==> (f a : f`A) = (a : A)" | |
| 274 | by (blast dest: injD) | |
| 275 | ||
| 276 | lemma inj_image_subset_iff: "inj f ==> (f`A <= f`B) = (A<=B)" | |
| 277 | by (simp add: inj_on_def, blast) | |
| 278 | ||
| 279 | lemma inj_image_eq_iff: "inj f ==> (f`A = f`B) = (A = B)" | |
| 280 | by (blast dest: injD) | |
| 281 | ||
| 282 | lemma image_UN: "(f ` (UNION A B)) = (UN x:A.(f ` (B x)))" | |
| 283 | by blast | |
| 284 | ||
| 285 | (*injectivity's required. Left-to-right inclusion holds even if A is empty*) | |
| 286 | lemma image_INT: | |
| 287 | "[| inj_on f C; ALL x:A. B x <= C; j:A |] | |
| 288 | ==> f ` (INTER A B) = (INT x:A. f ` B x)" | |
| 289 | apply (simp add: inj_on_def, blast) | |
| 290 | done | |
| 291 | ||
| 292 | (*Compare with image_INT: no use of inj_on, and if f is surjective then | |
| 293 | it doesn't matter whether A is empty*) | |
| 294 | lemma bij_image_INT: "bij f ==> f ` (INTER A B) = (INT x:A. f ` B x)" | |
| 295 | apply (simp add: bij_def) | |
| 296 | apply (simp add: inj_on_def surj_def, blast) | |
| 297 | done | |
| 298 | ||
| 299 | lemma surj_Compl_image_subset: "surj f ==> -(f`A) <= f`(-A)" | |
| 300 | by (auto simp add: surj_def) | |
| 301 | ||
| 302 | lemma inj_image_Compl_subset: "inj f ==> f`(-A) <= -(f`A)" | |
| 303 | by (auto simp add: inj_on_def) | |
| 5852 | 304 | |
| 13585 | 305 | lemma bij_image_Compl_eq: "bij f ==> f`(-A) = -(f`A)" | 
| 306 | apply (simp add: bij_def) | |
| 307 | apply (rule equalityI) | |
| 308 | apply (simp_all (no_asm_simp) add: inj_image_Compl_subset surj_Compl_image_subset) | |
| 309 | done | |
| 310 | ||
| 311 | ||
| 312 | subsection{*Function Updating*}
 | |
| 313 | ||
| 314 | lemma fun_upd_idem_iff: "(f(x:=y) = f) = (f x = y)" | |
| 315 | apply (simp add: fun_upd_def, safe) | |
| 316 | apply (erule subst) | |
| 317 | apply (rule_tac [2] ext, auto) | |
| 318 | done | |
| 319 | ||
| 320 | (* f x = y ==> f(x:=y) = f *) | |
| 321 | lemmas fun_upd_idem = fun_upd_idem_iff [THEN iffD2, standard] | |
| 322 | ||
| 323 | (* f(x := f x) = f *) | |
| 324 | declare refl [THEN fun_upd_idem, iff] | |
| 325 | ||
| 326 | lemma fun_upd_apply [simp]: "(f(x:=y))z = (if z=x then y else f z)" | |
| 327 | apply (simp (no_asm) add: fun_upd_def) | |
| 328 | done | |
| 329 | ||
| 330 | (* fun_upd_apply supersedes these two, but they are useful | |
| 331 | if fun_upd_apply is intentionally removed from the simpset *) | |
| 332 | lemma fun_upd_same: "(f(x:=y)) x = y" | |
| 333 | by simp | |
| 334 | ||
| 335 | lemma fun_upd_other: "z~=x ==> (f(x:=y)) z = f z" | |
| 336 | by simp | |
| 337 | ||
| 338 | lemma fun_upd_upd [simp]: "f(x:=y,x:=z) = f(x:=z)" | |
| 339 | by (simp add: expand_fun_eq) | |
| 340 | ||
| 341 | lemma fun_upd_twist: "a ~= c ==> (m(a:=b))(c:=d) = (m(c:=d))(a:=b)" | |
| 342 | by (rule ext, auto) | |
| 343 | ||
| 13910 | 344 | subsection{* overwrite *}
 | 
| 345 | ||
| 346 | lemma overwrite_emptyset[simp]: "f(g|{}) = f"
 | |
| 347 | by(simp add:overwrite_def) | |
| 348 | ||
| 349 | lemma overwrite_apply_notin[simp]: "a ~: A ==> (f(g|A)) a = f a" | |
| 350 | by(simp add:overwrite_def) | |
| 351 | ||
| 352 | lemma overwrite_apply_in[simp]: "a : A ==> (f(g|A)) a = g a" | |
| 353 | by(simp add:overwrite_def) | |
| 354 | ||
| 13585 | 355 | text{*The ML section includes some compatibility bindings and a simproc
 | 
| 356 | for function updates, in addition to the usual ML-bindings of theorems.*} | |
| 357 | ML | |
| 358 | {*
 | |
| 359 | val id_def = thm "id_def"; | |
| 360 | val inj_on_def = thm "inj_on_def"; | |
| 361 | val surj_def = thm "surj_def"; | |
| 362 | val bij_def = thm "bij_def"; | |
| 363 | val fun_upd_def = thm "fun_upd_def"; | |
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11123diff
changeset | 364 | |
| 13585 | 365 | val o_def = thm "comp_def"; | 
| 366 | val injI = thm "inj_onI"; | |
| 367 | val inj_inverseI = thm "inj_on_inverseI"; | |
| 368 | val set_cs = claset() delrules [equalityI]; | |
| 369 | ||
| 370 | val print_translation = [("Pi", dependent_tr' ("@Pi", "op funcset"))];
 | |
| 371 | ||
| 372 | (* simplifies terms of the form f(...,x:=y,...,x:=z,...) to f(...,x:=z,...) *) | |
| 373 | local | |
| 374 | fun gen_fun_upd None T _ _ = None | |
| 375 |     | gen_fun_upd (Some f) T x y = Some (Const ("Fun.fun_upd",T) $ f $ x $ y)
 | |
| 376 | fun dest_fun_T1 (Type (_, T :: Ts)) = T | |
| 377 |   fun find_double (t as Const ("Fun.fun_upd",T) $ f $ x $ y) =
 | |
| 378 | let | |
| 379 |       fun find (Const ("Fun.fun_upd",T) $ g $ v $ w) =
 | |
| 380 | if v aconv x then Some g else gen_fun_upd (find g) T v w | |
| 381 | | find t = None | |
| 382 | in (dest_fun_T1 T, gen_fun_upd (find f) T x y) end | |
| 383 | ||
| 384 | val ss = simpset () | |
| 385 | val fun_upd_prover = K (rtac eq_reflection 1 THEN rtac ext 1 THEN simp_tac ss 1) | |
| 386 | in | |
| 387 | val fun_upd2_simproc = | |
| 388 | Simplifier.simproc (Theory.sign_of (the_context ())) | |
| 389 | "fun_upd2" ["f(v := w, x := y)"] | |
| 390 | (fn sg => fn _ => fn t => | |
| 391 | case find_double t of (T, None) => None | |
| 392 | | (T, Some rhs) => Some (Tactic.prove sg [] [] (Term.equals T $ t $ rhs) fun_upd_prover)) | |
| 393 | end; | |
| 394 | Addsimprocs[fun_upd2_simproc]; | |
| 5852 | 395 | |
| 13585 | 396 | val expand_fun_eq = thm "expand_fun_eq"; | 
| 397 | val apply_inverse = thm "apply_inverse"; | |
| 398 | val id_apply = thm "id_apply"; | |
| 399 | val o_apply = thm "o_apply"; | |
| 400 | val o_assoc = thm "o_assoc"; | |
| 401 | val id_o = thm "id_o"; | |
| 402 | val o_id = thm "o_id"; | |
| 403 | val image_compose = thm "image_compose"; | |
| 404 | val image_eq_UN = thm "image_eq_UN"; | |
| 405 | val UN_o = thm "UN_o"; | |
| 406 | val datatype_injI = thm "datatype_injI"; | |
| 407 | val injD = thm "injD"; | |
| 408 | val inj_eq = thm "inj_eq"; | |
| 409 | val inj_onI = thm "inj_onI"; | |
| 410 | val inj_on_inverseI = thm "inj_on_inverseI"; | |
| 411 | val inj_onD = thm "inj_onD"; | |
| 412 | val inj_on_iff = thm "inj_on_iff"; | |
| 413 | val comp_inj_on = thm "comp_inj_on"; | |
| 414 | val inj_on_contraD = thm "inj_on_contraD"; | |
| 415 | val inj_singleton = thm "inj_singleton"; | |
| 416 | val subset_inj_on = thm "subset_inj_on"; | |
| 417 | val surjI = thm "surjI"; | |
| 418 | val surj_range = thm "surj_range"; | |
| 419 | val surjD = thm "surjD"; | |
| 420 | val surjE = thm "surjE"; | |
| 421 | val comp_surj = thm "comp_surj"; | |
| 422 | val bijI = thm "bijI"; | |
| 423 | val bij_is_inj = thm "bij_is_inj"; | |
| 424 | val bij_is_surj = thm "bij_is_surj"; | |
| 425 | val image_ident = thm "image_ident"; | |
| 426 | val image_id = thm "image_id"; | |
| 427 | val vimage_ident = thm "vimage_ident"; | |
| 428 | val vimage_id = thm "vimage_id"; | |
| 429 | val vimage_image_eq = thm "vimage_image_eq"; | |
| 430 | val image_vimage_subset = thm "image_vimage_subset"; | |
| 431 | val image_vimage_eq = thm "image_vimage_eq"; | |
| 432 | val surj_image_vimage_eq = thm "surj_image_vimage_eq"; | |
| 433 | val inj_vimage_image_eq = thm "inj_vimage_image_eq"; | |
| 434 | val vimage_subsetD = thm "vimage_subsetD"; | |
| 435 | val vimage_subsetI = thm "vimage_subsetI"; | |
| 436 | val vimage_subset_eq = thm "vimage_subset_eq"; | |
| 437 | val image_Int_subset = thm "image_Int_subset"; | |
| 438 | val image_diff_subset = thm "image_diff_subset"; | |
| 439 | val inj_on_image_Int = thm "inj_on_image_Int"; | |
| 440 | val inj_on_image_set_diff = thm "inj_on_image_set_diff"; | |
| 441 | val image_Int = thm "image_Int"; | |
| 442 | val image_set_diff = thm "image_set_diff"; | |
| 443 | val inj_image_mem_iff = thm "inj_image_mem_iff"; | |
| 444 | val inj_image_subset_iff = thm "inj_image_subset_iff"; | |
| 445 | val inj_image_eq_iff = thm "inj_image_eq_iff"; | |
| 446 | val image_UN = thm "image_UN"; | |
| 447 | val image_INT = thm "image_INT"; | |
| 448 | val bij_image_INT = thm "bij_image_INT"; | |
| 449 | val surj_Compl_image_subset = thm "surj_Compl_image_subset"; | |
| 450 | val inj_image_Compl_subset = thm "inj_image_Compl_subset"; | |
| 451 | val bij_image_Compl_eq = thm "bij_image_Compl_eq"; | |
| 452 | val fun_upd_idem_iff = thm "fun_upd_idem_iff"; | |
| 453 | val fun_upd_idem = thm "fun_upd_idem"; | |
| 454 | val fun_upd_apply = thm "fun_upd_apply"; | |
| 455 | val fun_upd_same = thm "fun_upd_same"; | |
| 456 | val fun_upd_other = thm "fun_upd_other"; | |
| 457 | val fun_upd_upd = thm "fun_upd_upd"; | |
| 458 | val fun_upd_twist = thm "fun_upd_twist"; | |
| 13637 | 459 | val range_ex1_eq = thm "range_ex1_eq"; | 
| 13585 | 460 | *} | 
| 5852 | 461 | |
| 2912 | 462 | end |