| author | wenzelm | 
| Thu, 10 May 2007 00:39:51 +0200 | |
| changeset 22905 | dab6a898b47c | 
| parent 22886 | cdff6ef76009 | 
| child 23738 | 3a801ffdc58c | 
| 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 | 
| 18154 | 5 | *) | 
| 923 | 6 | |
| 18154 | 7 | header {* Notions about functions *}
 | 
| 923 | 8 | |
| 15510 | 9 | theory Fun | 
| 22886 | 10 | imports Set | 
| 15131 | 11 | begin | 
| 2912 | 12 | |
| 13585 | 13 | constdefs | 
| 14 |   fun_upd :: "('a => 'b) => 'a => 'b => ('a => 'b)"
 | |
| 22886 | 15 | "fun_upd f a b == % x. if x=a then b else f x" | 
| 6171 | 16 | |
| 9141 | 17 | nonterminals | 
| 18 | updbinds updbind | |
| 5305 | 19 | syntax | 
| 13585 | 20 |   "_updbind" :: "['a, 'a] => updbind"             ("(2_ :=/ _)")
 | 
| 21 |   ""         :: "updbind => updbinds"             ("_")
 | |
| 22 |   "_updbinds":: "[updbind, updbinds] => updbinds" ("_,/ _")
 | |
| 23 |   "_Update"  :: "['a, updbinds] => 'a"            ("_/'((_)')" [1000,0] 900)
 | |
| 5305 | 24 | |
| 25 | translations | |
| 26 | "_Update f (_updbinds b bs)" == "_Update (_Update f b) bs" | |
| 27 | "f(x:=y)" == "fun_upd f x y" | |
| 2912 | 28 | |
| 9340 | 29 | (* Hint: to define the sum of two functions (or maps), use sum_case. | 
| 30 | A nice infix syntax could be defined (in Datatype.thy or below) by | |
| 31 | consts | |
| 32 |   fun_sum :: "('a => 'c) => ('b => 'c) => (('a+'b) => 'c)" (infixr "'(+')"80)
 | |
| 33 | translations | |
| 13585 | 34 | "fun_sum" == sum_case | 
| 9340 | 35 | *) | 
| 12258 | 36 | |
| 22744 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 37 | definition | 
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 38 |   override_on :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> 'b"
 | 
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 39 | where | 
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 40 | "override_on f g A = (\<lambda>a. if a \<in> A then g a else f a)" | 
| 6171 | 41 | |
| 22744 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 42 | definition | 
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 43 | id :: "'a \<Rightarrow> 'a" | 
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 44 | where | 
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 45 | "id = (\<lambda>x. x)" | 
| 13910 | 46 | |
| 22744 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 47 | definition | 
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 48 |   comp :: "('b \<Rightarrow> 'c) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'c" (infixl "o" 55)
 | 
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 49 | where | 
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 50 | "f o g = (\<lambda>x. f (g x))" | 
| 11123 | 51 | |
| 21210 | 52 | notation (xsymbols) | 
| 19656 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19536diff
changeset | 53 | comp (infixl "\<circ>" 55) | 
| 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19536diff
changeset | 54 | |
| 21210 | 55 | notation (HTML output) | 
| 19656 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19536diff
changeset | 56 | comp (infixl "\<circ>" 55) | 
| 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19536diff
changeset | 57 | |
| 13585 | 58 | text{*compatibility*}
 | 
| 59 | lemmas o_def = comp_def | |
| 2912 | 60 | |
| 13585 | 61 | constdefs | 
| 62 | inj_on :: "['a => 'b, 'a set] => bool" (*injective*) | |
| 19363 | 63 | "inj_on f A == ! x:A. ! y:A. f(x)=f(y) --> x=y" | 
| 6171 | 64 | |
| 13585 | 65 | text{*A common special case: functions injective over the entire domain type.*}
 | 
| 19323 | 66 | |
| 19363 | 67 | abbreviation | 
| 68 | "inj f == inj_on f UNIV" | |
| 5852 | 69 | |
| 7374 | 70 | constdefs | 
| 13585 | 71 |   surj :: "('a => 'b) => bool"                   (*surjective*)
 | 
| 19363 | 72 | "surj f == ! y. ? x. y=f(x)" | 
| 12258 | 73 | |
| 13585 | 74 |   bij :: "('a => 'b) => bool"                    (*bijective*)
 | 
| 19363 | 75 | "bij f == inj f & surj f" | 
| 12258 | 76 | |
| 7374 | 77 | |
| 13585 | 78 | |
| 79 | text{*As a simplification rule, it replaces all function equalities by
 | |
| 80 | first-order equalities.*} | |
| 21327 | 81 | lemma expand_fun_eq: "f = g \<longleftrightarrow> (\<forall>x. f x = g x)" | 
| 13585 | 82 | apply (rule iffI) | 
| 83 | apply (simp (no_asm_simp)) | |
| 21327 | 84 | apply (rule ext) | 
| 85 | apply (simp (no_asm_simp)) | |
| 13585 | 86 | done | 
| 87 | ||
| 88 | lemma apply_inverse: | |
| 89 | "[| f(x)=u; !!x. P(x) ==> g(f(x)) = x; P(x) |] ==> x=g(u)" | |
| 90 | by auto | |
| 91 | ||
| 92 | ||
| 93 | text{*The Identity Function: @{term id}*}
 | |
| 94 | lemma id_apply [simp]: "id x = x" | |
| 95 | by (simp add: id_def) | |
| 96 | ||
| 16733 
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
 nipkow parents: 
15691diff
changeset | 97 | lemma inj_on_id[simp]: "inj_on id A" | 
| 15510 | 98 | by (simp add: inj_on_def) | 
| 99 | ||
| 16733 
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
 nipkow parents: 
15691diff
changeset | 100 | lemma inj_on_id2[simp]: "inj_on (%x. x) A" | 
| 
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
 nipkow parents: 
15691diff
changeset | 101 | by (simp add: inj_on_def) | 
| 
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
 nipkow parents: 
15691diff
changeset | 102 | |
| 
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
 nipkow parents: 
15691diff
changeset | 103 | lemma surj_id[simp]: "surj id" | 
| 15510 | 104 | by (simp add: surj_def) | 
| 105 | ||
| 16733 
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
 nipkow parents: 
15691diff
changeset | 106 | lemma bij_id[simp]: "bij id" | 
| 15510 | 107 | by (simp add: bij_def inj_on_id surj_id) | 
| 108 | ||
| 109 | ||
| 13585 | 110 | |
| 111 | subsection{*The Composition Operator: @{term "f \<circ> g"}*}
 | |
| 112 | ||
| 113 | lemma o_apply [simp]: "(f o g) x = f (g x)" | |
| 114 | by (simp add: comp_def) | |
| 115 | ||
| 116 | lemma o_assoc: "f o (g o h) = f o g o h" | |
| 117 | by (simp add: comp_def) | |
| 118 | ||
| 119 | lemma id_o [simp]: "id o g = g" | |
| 120 | by (simp add: comp_def) | |
| 121 | ||
| 122 | lemma o_id [simp]: "f o id = f" | |
| 123 | by (simp add: comp_def) | |
| 124 | ||
| 125 | lemma image_compose: "(f o g) ` r = f`(g`r)" | |
| 126 | by (simp add: comp_def, blast) | |
| 127 | ||
| 128 | lemma image_eq_UN: "f`A = (UN x:A. {f x})"
 | |
| 129 | by blast | |
| 130 | ||
| 131 | lemma UN_o: "UNION A (g o f) = UNION (f`A) g" | |
| 132 | by (unfold comp_def, blast) | |
| 133 | ||
| 134 | ||
| 135 | subsection{*The Injectivity Predicate, @{term inj}*}
 | |
| 136 | ||
| 137 | text{*NB: @{term inj} now just translates to @{term inj_on}*}
 | |
| 138 | ||
| 139 | ||
| 140 | text{*For Proofs in @{text "Tools/datatype_rep_proofs"}*}
 | |
| 141 | lemma datatype_injI: | |
| 142 | "(!! x. ALL y. f(x) = f(y) --> x=y) ==> inj(f)" | |
| 143 | by (simp add: inj_on_def) | |
| 144 | ||
| 13637 | 145 | theorem range_ex1_eq: "inj f \<Longrightarrow> b : range f = (EX! x. b = f x)" | 
| 146 | by (unfold inj_on_def, blast) | |
| 147 | ||
| 13585 | 148 | lemma injD: "[| inj(f); f(x) = f(y) |] ==> x=y" | 
| 149 | by (simp add: inj_on_def) | |
| 150 | ||
| 151 | (*Useful with the simplifier*) | |
| 152 | lemma inj_eq: "inj(f) ==> (f(x) = f(y)) = (x=y)" | |
| 153 | by (force simp add: inj_on_def) | |
| 154 | ||
| 155 | ||
| 156 | subsection{*The Predicate @{term inj_on}: Injectivity On A Restricted Domain*}
 | |
| 157 | ||
| 158 | lemma inj_onI: | |
| 159 | "(!! x y. [| x:A; y:A; f(x) = f(y) |] ==> x=y) ==> inj_on f A" | |
| 160 | by (simp add: inj_on_def) | |
| 161 | ||
| 162 | lemma inj_on_inverseI: "(!!x. x:A ==> g(f(x)) = x) ==> inj_on f A" | |
| 163 | by (auto dest: arg_cong [of concl: g] simp add: inj_on_def) | |
| 164 | ||
| 165 | lemma inj_onD: "[| inj_on f A; f(x)=f(y); x:A; y:A |] ==> x=y" | |
| 166 | by (unfold inj_on_def, blast) | |
| 167 | ||
| 168 | lemma inj_on_iff: "[| inj_on f A; x:A; y:A |] ==> (f(x)=f(y)) = (x=y)" | |
| 169 | by (blast dest!: inj_onD) | |
| 170 | ||
| 171 | lemma comp_inj_on: | |
| 172 | "[| inj_on f A; inj_on g (f`A) |] ==> inj_on (g o f) A" | |
| 173 | by (simp add: comp_def inj_on_def) | |
| 174 | ||
| 15303 | 175 | lemma inj_on_imageI: "inj_on (g o f) A \<Longrightarrow> inj_on g (f ` A)" | 
| 176 | apply(simp add:inj_on_def image_def) | |
| 177 | apply blast | |
| 178 | done | |
| 179 | ||
| 15439 | 180 | lemma inj_on_image_iff: "\<lbrakk> ALL x:A. ALL y:A. (g(f x) = g(f y)) = (g x = g y); | 
| 181 | inj_on f A \<rbrakk> \<Longrightarrow> inj_on g (f ` A) = inj_on g A" | |
| 182 | apply(unfold inj_on_def) | |
| 183 | apply blast | |
| 184 | done | |
| 185 | ||
| 13585 | 186 | lemma inj_on_contraD: "[| inj_on f A; ~x=y; x:A; y:A |] ==> ~ f(x)=f(y)" | 
| 187 | by (unfold inj_on_def, blast) | |
| 12258 | 188 | |
| 13585 | 189 | lemma inj_singleton: "inj (%s. {s})"
 | 
| 190 | by (simp add: inj_on_def) | |
| 191 | ||
| 15111 | 192 | lemma inj_on_empty[iff]: "inj_on f {}"
 | 
| 193 | by(simp add: inj_on_def) | |
| 194 | ||
| 15303 | 195 | lemma subset_inj_on: "[| inj_on f B; A <= B |] ==> inj_on f A" | 
| 13585 | 196 | by (unfold inj_on_def, blast) | 
| 197 | ||
| 15111 | 198 | lemma inj_on_Un: | 
| 199 | "inj_on f (A Un B) = | |
| 200 |   (inj_on f A & inj_on f B & f`(A-B) Int f`(B-A) = {})"
 | |
| 201 | apply(unfold inj_on_def) | |
| 202 | apply (blast intro:sym) | |
| 203 | done | |
| 204 | ||
| 205 | lemma inj_on_insert[iff]: | |
| 206 |   "inj_on f (insert a A) = (inj_on f A & f a ~: f`(A-{a}))"
 | |
| 207 | apply(unfold inj_on_def) | |
| 208 | apply (blast intro:sym) | |
| 209 | done | |
| 210 | ||
| 211 | lemma inj_on_diff: "inj_on f A ==> inj_on f (A-B)" | |
| 212 | apply(unfold inj_on_def) | |
| 213 | apply (blast) | |
| 214 | done | |
| 215 | ||
| 13585 | 216 | |
| 217 | subsection{*The Predicate @{term surj}: Surjectivity*}
 | |
| 218 | ||
| 219 | lemma surjI: "(!! x. g(f x) = x) ==> surj g" | |
| 220 | apply (simp add: surj_def) | |
| 221 | apply (blast intro: sym) | |
| 222 | done | |
| 223 | ||
| 224 | lemma surj_range: "surj f ==> range f = UNIV" | |
| 225 | by (auto simp add: surj_def) | |
| 226 | ||
| 227 | lemma surjD: "surj f ==> EX x. y = f x" | |
| 228 | by (simp add: surj_def) | |
| 229 | ||
| 230 | lemma surjE: "surj f ==> (!!x. y = f x ==> C) ==> C" | |
| 231 | by (simp add: surj_def, blast) | |
| 232 | ||
| 233 | lemma comp_surj: "[| surj f; surj g |] ==> surj (g o f)" | |
| 234 | apply (simp add: comp_def surj_def, clarify) | |
| 235 | apply (drule_tac x = y in spec, clarify) | |
| 236 | apply (drule_tac x = x in spec, blast) | |
| 237 | done | |
| 238 | ||
| 239 | ||
| 240 | ||
| 241 | subsection{*The Predicate @{term bij}: Bijectivity*}
 | |
| 242 | ||
| 243 | lemma bijI: "[| inj f; surj f |] ==> bij f" | |
| 244 | by (simp add: bij_def) | |
| 245 | ||
| 246 | lemma bij_is_inj: "bij f ==> inj f" | |
| 247 | by (simp add: bij_def) | |
| 248 | ||
| 249 | lemma bij_is_surj: "bij f ==> surj f" | |
| 250 | by (simp add: bij_def) | |
| 251 | ||
| 252 | ||
| 253 | subsection{*Facts About the Identity Function*}
 | |
| 5852 | 254 | |
| 13585 | 255 | text{*We seem to need both the @{term id} forms and the @{term "\<lambda>x. x"}
 | 
| 256 | forms. The latter can arise by rewriting, while @{term id} may be used
 | |
| 257 | explicitly.*} | |
| 258 | ||
| 259 | lemma image_ident [simp]: "(%x. x) ` Y = Y" | |
| 260 | by blast | |
| 261 | ||
| 262 | lemma image_id [simp]: "id ` Y = Y" | |
| 263 | by (simp add: id_def) | |
| 264 | ||
| 265 | lemma vimage_ident [simp]: "(%x. x) -` Y = Y" | |
| 266 | by blast | |
| 267 | ||
| 268 | lemma vimage_id [simp]: "id -` A = A" | |
| 269 | by (simp add: id_def) | |
| 270 | ||
| 271 | lemma vimage_image_eq: "f -` (f ` A) = {y. EX x:A. f x = f y}"
 | |
| 272 | by (blast intro: sym) | |
| 273 | ||
| 274 | lemma image_vimage_subset: "f ` (f -` A) <= A" | |
| 275 | by blast | |
| 276 | ||
| 277 | lemma image_vimage_eq [simp]: "f ` (f -` A) = A Int range f" | |
| 278 | by blast | |
| 279 | ||
| 280 | lemma surj_image_vimage_eq: "surj f ==> f ` (f -` A) = A" | |
| 281 | by (simp add: surj_range) | |
| 282 | ||
| 283 | lemma inj_vimage_image_eq: "inj f ==> f -` (f ` A) = A" | |
| 284 | by (simp add: inj_on_def, blast) | |
| 285 | ||
| 286 | lemma vimage_subsetD: "surj f ==> f -` B <= A ==> B <= f ` A" | |
| 287 | apply (unfold surj_def) | |
| 288 | apply (blast intro: sym) | |
| 289 | done | |
| 290 | ||
| 291 | lemma vimage_subsetI: "inj f ==> B <= f ` A ==> f -` B <= A" | |
| 292 | by (unfold inj_on_def, blast) | |
| 293 | ||
| 294 | lemma vimage_subset_eq: "bij f ==> (f -` B <= A) = (B <= f ` A)" | |
| 295 | apply (unfold bij_def) | |
| 296 | apply (blast del: subsetI intro: vimage_subsetI vimage_subsetD) | |
| 297 | done | |
| 298 | ||
| 299 | lemma image_Int_subset: "f`(A Int B) <= f`A Int f`B" | |
| 300 | by blast | |
| 301 | ||
| 302 | lemma image_diff_subset: "f`A - f`B <= f`(A - B)" | |
| 303 | by blast | |
| 5852 | 304 | |
| 13585 | 305 | lemma inj_on_image_Int: | 
| 306 | "[| inj_on f C; A<=C; B<=C |] ==> f`(A Int B) = f`A Int f`B" | |
| 307 | apply (simp add: inj_on_def, blast) | |
| 308 | done | |
| 309 | ||
| 310 | lemma inj_on_image_set_diff: | |
| 311 | "[| inj_on f C; A<=C; B<=C |] ==> f`(A-B) = f`A - f`B" | |
| 312 | apply (simp add: inj_on_def, blast) | |
| 313 | done | |
| 314 | ||
| 315 | lemma image_Int: "inj f ==> f`(A Int B) = f`A Int f`B" | |
| 316 | by (simp add: inj_on_def, blast) | |
| 317 | ||
| 318 | lemma image_set_diff: "inj f ==> f`(A-B) = f`A - f`B" | |
| 319 | by (simp add: inj_on_def, blast) | |
| 320 | ||
| 321 | lemma inj_image_mem_iff: "inj f ==> (f a : f`A) = (a : A)" | |
| 322 | by (blast dest: injD) | |
| 323 | ||
| 324 | lemma inj_image_subset_iff: "inj f ==> (f`A <= f`B) = (A<=B)" | |
| 325 | by (simp add: inj_on_def, blast) | |
| 326 | ||
| 327 | lemma inj_image_eq_iff: "inj f ==> (f`A = f`B) = (A = B)" | |
| 328 | by (blast dest: injD) | |
| 329 | ||
| 330 | lemma image_UN: "(f ` (UNION A B)) = (UN x:A.(f ` (B x)))" | |
| 331 | by blast | |
| 332 | ||
| 333 | (*injectivity's required. Left-to-right inclusion holds even if A is empty*) | |
| 334 | lemma image_INT: | |
| 335 | "[| inj_on f C; ALL x:A. B x <= C; j:A |] | |
| 336 | ==> f ` (INTER A B) = (INT x:A. f ` B x)" | |
| 337 | apply (simp add: inj_on_def, blast) | |
| 338 | done | |
| 339 | ||
| 340 | (*Compare with image_INT: no use of inj_on, and if f is surjective then | |
| 341 | it doesn't matter whether A is empty*) | |
| 342 | lemma bij_image_INT: "bij f ==> f ` (INTER A B) = (INT x:A. f ` B x)" | |
| 343 | apply (simp add: bij_def) | |
| 344 | apply (simp add: inj_on_def surj_def, blast) | |
| 345 | done | |
| 346 | ||
| 347 | lemma surj_Compl_image_subset: "surj f ==> -(f`A) <= f`(-A)" | |
| 348 | by (auto simp add: surj_def) | |
| 349 | ||
| 350 | lemma inj_image_Compl_subset: "inj f ==> f`(-A) <= -(f`A)" | |
| 351 | by (auto simp add: inj_on_def) | |
| 5852 | 352 | |
| 13585 | 353 | lemma bij_image_Compl_eq: "bij f ==> f`(-A) = -(f`A)" | 
| 354 | apply (simp add: bij_def) | |
| 355 | apply (rule equalityI) | |
| 356 | apply (simp_all (no_asm_simp) add: inj_image_Compl_subset surj_Compl_image_subset) | |
| 357 | done | |
| 358 | ||
| 359 | ||
| 360 | subsection{*Function Updating*}
 | |
| 361 | ||
| 362 | lemma fun_upd_idem_iff: "(f(x:=y) = f) = (f x = y)" | |
| 363 | apply (simp add: fun_upd_def, safe) | |
| 364 | apply (erule subst) | |
| 365 | apply (rule_tac [2] ext, auto) | |
| 366 | done | |
| 367 | ||
| 368 | (* f x = y ==> f(x:=y) = f *) | |
| 369 | lemmas fun_upd_idem = fun_upd_idem_iff [THEN iffD2, standard] | |
| 370 | ||
| 371 | (* f(x := f x) = f *) | |
| 17084 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16973diff
changeset | 372 | lemmas fun_upd_triv = refl [THEN fun_upd_idem] | 
| 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16973diff
changeset | 373 | declare fun_upd_triv [iff] | 
| 13585 | 374 | |
| 375 | lemma fun_upd_apply [simp]: "(f(x:=y))z = (if z=x then y else f z)" | |
| 17084 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16973diff
changeset | 376 | by (simp add: fun_upd_def) | 
| 13585 | 377 | |
| 378 | (* fun_upd_apply supersedes these two, but they are useful | |
| 379 | if fun_upd_apply is intentionally removed from the simpset *) | |
| 380 | lemma fun_upd_same: "(f(x:=y)) x = y" | |
| 381 | by simp | |
| 382 | ||
| 383 | lemma fun_upd_other: "z~=x ==> (f(x:=y)) z = f z" | |
| 384 | by simp | |
| 385 | ||
| 386 | lemma fun_upd_upd [simp]: "f(x:=y,x:=z) = f(x:=z)" | |
| 387 | by (simp add: expand_fun_eq) | |
| 388 | ||
| 389 | lemma fun_upd_twist: "a ~= c ==> (m(a:=b))(c:=d) = (m(c:=d))(a:=b)" | |
| 390 | by (rule ext, auto) | |
| 391 | ||
| 15303 | 392 | lemma inj_on_fun_updI: "\<lbrakk> inj_on f A; y \<notin> f`A \<rbrakk> \<Longrightarrow> inj_on (f(x:=y)) A" | 
| 393 | by(fastsimp simp:inj_on_def image_def) | |
| 394 | ||
| 15510 | 395 | lemma fun_upd_image: | 
| 396 |      "f(x:=y) ` A = (if x \<in> A then insert y (f ` (A-{x})) else f ` A)"
 | |
| 397 | by auto | |
| 398 | ||
| 15691 | 399 | subsection{* @{text override_on} *}
 | 
| 13910 | 400 | |
| 15691 | 401 | lemma override_on_emptyset[simp]: "override_on f g {} = f"
 | 
| 402 | by(simp add:override_on_def) | |
| 13910 | 403 | |
| 15691 | 404 | lemma override_on_apply_notin[simp]: "a ~: A ==> (override_on f g A) a = f a" | 
| 405 | by(simp add:override_on_def) | |
| 13910 | 406 | |
| 15691 | 407 | lemma override_on_apply_in[simp]: "a : A ==> (override_on f g A) a = g a" | 
| 408 | by(simp add:override_on_def) | |
| 13910 | 409 | |
| 15510 | 410 | subsection{* swap *}
 | 
| 411 | ||
| 22744 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 412 | definition | 
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 413 |   swap :: "'a \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b)"
 | 
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 414 | where | 
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 415 | "swap a b f = f (a := f b, b:= f a)" | 
| 15510 | 416 | |
| 417 | lemma swap_self: "swap a a f = f" | |
| 15691 | 418 | by (simp add: swap_def) | 
| 15510 | 419 | |
| 420 | lemma swap_commute: "swap a b f = swap b a f" | |
| 421 | by (rule ext, simp add: fun_upd_def swap_def) | |
| 422 | ||
| 423 | lemma swap_nilpotent [simp]: "swap a b (swap a b f) = f" | |
| 424 | by (rule ext, simp add: fun_upd_def swap_def) | |
| 425 | ||
| 426 | lemma inj_on_imp_inj_on_swap: | |
| 22744 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 427 | "[|inj_on f A; a \<in> A; b \<in> A|] ==> inj_on (swap a b f) A" | 
| 15510 | 428 | by (simp add: inj_on_def swap_def, blast) | 
| 429 | ||
| 430 | lemma inj_on_swap_iff [simp]: | |
| 431 | assumes A: "a \<in> A" "b \<in> A" shows "inj_on (swap a b f) A = inj_on f A" | |
| 432 | proof | |
| 433 | assume "inj_on (swap a b f) A" | |
| 434 | with A have "inj_on (swap a b (swap a b f)) A" | |
| 17589 | 435 | by (iprover intro: inj_on_imp_inj_on_swap) | 
| 15510 | 436 | thus "inj_on f A" by simp | 
| 437 | next | |
| 438 | assume "inj_on f A" | |
| 17589 | 439 | with A show "inj_on (swap a b f) A" by (iprover intro: inj_on_imp_inj_on_swap) | 
| 15510 | 440 | qed | 
| 441 | ||
| 442 | lemma surj_imp_surj_swap: "surj f ==> surj (swap a b f)" | |
| 443 | apply (simp add: surj_def swap_def, clarify) | |
| 444 | apply (rule_tac P = "y = f b" in case_split_thm, blast) | |
| 445 | apply (rule_tac P = "y = f a" in case_split_thm, auto) | |
| 446 |   --{*We don't yet have @{text case_tac}*}
 | |
| 447 | done | |
| 448 | ||
| 449 | lemma surj_swap_iff [simp]: "surj (swap a b f) = surj f" | |
| 450 | proof | |
| 451 | assume "surj (swap a b f)" | |
| 452 | hence "surj (swap a b (swap a b f))" by (rule surj_imp_surj_swap) | |
| 453 | thus "surj f" by simp | |
| 454 | next | |
| 455 | assume "surj f" | |
| 456 | thus "surj (swap a b f)" by (rule surj_imp_surj_swap) | |
| 457 | qed | |
| 458 | ||
| 459 | lemma bij_swap_iff: "bij (swap a b f) = bij f" | |
| 460 | by (simp add: bij_def) | |
| 21547 
9c9fdf4c2949
moved order arities for fun and bool to Fun/Orderings
 haftmann parents: 
21327diff
changeset | 461 | |
| 
9c9fdf4c2949
moved order arities for fun and bool to Fun/Orderings
 haftmann parents: 
21327diff
changeset | 462 | |
| 22453 | 463 | subsection {* Order and lattice on functions *}
 | 
| 464 | ||
| 465 | instance "fun" :: (type, ord) ord | |
| 466 | le_fun_def: "f \<le> g \<equiv> \<forall>x. f x \<le> g x" | |
| 467 | less_fun_def: "f < g \<equiv> f \<le> g \<and> f \<noteq> g" .. | |
| 468 | ||
| 22845 | 469 | lemmas [code func del] = le_fun_def less_fun_def | 
| 21547 
9c9fdf4c2949
moved order arities for fun and bool to Fun/Orderings
 haftmann parents: 
21327diff
changeset | 470 | |
| 
9c9fdf4c2949
moved order arities for fun and bool to Fun/Orderings
 haftmann parents: 
21327diff
changeset | 471 | instance "fun" :: (type, order) order | 
| 22845 | 472 | by default | 
| 473 | (auto simp add: le_fun_def less_fun_def expand_fun_eq | |
| 474 | intro: order_trans order_antisym) | |
| 21547 
9c9fdf4c2949
moved order arities for fun and bool to Fun/Orderings
 haftmann parents: 
21327diff
changeset | 475 | |
| 
9c9fdf4c2949
moved order arities for fun and bool to Fun/Orderings
 haftmann parents: 
21327diff
changeset | 476 | lemma le_funI: "(\<And>x. f x \<le> g x) \<Longrightarrow> f \<le> g" | 
| 
9c9fdf4c2949
moved order arities for fun and bool to Fun/Orderings
 haftmann parents: 
21327diff
changeset | 477 | unfolding le_fun_def by simp | 
| 
9c9fdf4c2949
moved order arities for fun and bool to Fun/Orderings
 haftmann parents: 
21327diff
changeset | 478 | |
| 
9c9fdf4c2949
moved order arities for fun and bool to Fun/Orderings
 haftmann parents: 
21327diff
changeset | 479 | lemma le_funE: "f \<le> g \<Longrightarrow> (f x \<le> g x \<Longrightarrow> P) \<Longrightarrow> P" | 
| 
9c9fdf4c2949
moved order arities for fun and bool to Fun/Orderings
 haftmann parents: 
21327diff
changeset | 480 | unfolding le_fun_def by simp | 
| 
9c9fdf4c2949
moved order arities for fun and bool to Fun/Orderings
 haftmann parents: 
21327diff
changeset | 481 | |
| 
9c9fdf4c2949
moved order arities for fun and bool to Fun/Orderings
 haftmann parents: 
21327diff
changeset | 482 | lemma le_funD: "f \<le> g \<Longrightarrow> f x \<le> g x" | 
| 
9c9fdf4c2949
moved order arities for fun and bool to Fun/Orderings
 haftmann parents: 
21327diff
changeset | 483 | unfolding le_fun_def by simp | 
| 
9c9fdf4c2949
moved order arities for fun and bool to Fun/Orderings
 haftmann parents: 
21327diff
changeset | 484 | |
| 22453 | 485 | text {*
 | 
| 486 |   Handy introduction and elimination rules for @{text "\<le>"}
 | |
| 487 | on unary and binary predicates | |
| 488 | *} | |
| 489 | ||
| 490 | lemma predicate1I [Pure.intro!, intro!]: | |
| 491 | assumes PQ: "\<And>x. P x \<Longrightarrow> Q x" | |
| 492 | shows "P \<le> Q" | |
| 493 | apply (rule le_funI) | |
| 494 | apply (rule le_boolI) | |
| 495 | apply (rule PQ) | |
| 496 | apply assumption | |
| 497 | done | |
| 498 | ||
| 499 | lemma predicate1D [Pure.dest, dest]: "P \<le> Q \<Longrightarrow> P x \<Longrightarrow> Q x" | |
| 500 | apply (erule le_funE) | |
| 501 | apply (erule le_boolE) | |
| 502 | apply assumption+ | |
| 503 | done | |
| 504 | ||
| 505 | lemma predicate2I [Pure.intro!, intro!]: | |
| 506 | assumes PQ: "\<And>x y. P x y \<Longrightarrow> Q x y" | |
| 507 | shows "P \<le> Q" | |
| 508 | apply (rule le_funI)+ | |
| 509 | apply (rule le_boolI) | |
| 510 | apply (rule PQ) | |
| 511 | apply assumption | |
| 512 | done | |
| 513 | ||
| 514 | lemma predicate2D [Pure.dest, dest]: "P \<le> Q \<Longrightarrow> P x y \<Longrightarrow> Q x y" | |
| 515 | apply (erule le_funE)+ | |
| 516 | apply (erule le_boolE) | |
| 517 | apply assumption+ | |
| 518 | done | |
| 519 | ||
| 520 | lemma rev_predicate1D: "P x ==> P <= Q ==> Q x" | |
| 521 | by (rule predicate1D) | |
| 522 | ||
| 523 | lemma rev_predicate2D: "P x y ==> P <= Q ==> Q x y" | |
| 524 | by (rule predicate2D) | |
| 525 | ||
| 526 | instance "fun" :: (type, lattice) lattice | |
| 527 | inf_fun_eq: "inf f g \<equiv> (\<lambda>x. inf (f x) (g x))" | |
| 528 | sup_fun_eq: "sup f g \<equiv> (\<lambda>x. sup (f x) (g x))" | |
| 529 | apply intro_classes | |
| 530 | unfolding inf_fun_eq sup_fun_eq | |
| 531 | apply (auto intro: le_funI) | |
| 532 | apply (rule le_funI) | |
| 533 | apply (auto dest: le_funD) | |
| 534 | apply (rule le_funI) | |
| 535 | apply (auto dest: le_funD) | |
| 536 | done | |
| 537 | ||
| 22845 | 538 | lemmas [code func del] = inf_fun_eq sup_fun_eq | 
| 22744 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22577diff
changeset | 539 | |
| 22453 | 540 | instance "fun" :: (type, distrib_lattice) distrib_lattice | 
| 541 | by default (auto simp add: inf_fun_eq sup_fun_eq sup_inf_distrib1) | |
| 21547 
9c9fdf4c2949
moved order arities for fun and bool to Fun/Orderings
 haftmann parents: 
21327diff
changeset | 542 | |
| 
9c9fdf4c2949
moved order arities for fun and bool to Fun/Orderings
 haftmann parents: 
21327diff
changeset | 543 | |
| 22845 | 544 | subsection {* Proof tool setup *} 
 | 
| 545 | ||
| 546 | text {* simplifies terms of the form
 | |
| 547 | f(...,x:=y,...,x:=z,...) to f(...,x:=z,...) *} | |
| 548 | ||
| 549 | ML {*
 | |
| 550 | let | |
| 551 | fun gen_fun_upd NONE T _ _ = NONE | |
| 552 |     | gen_fun_upd (SOME f) T x y = SOME (Const (@{const_name fun_upd},T) $ f $ x $ y)
 | |
| 553 | fun dest_fun_T1 (Type (_, T :: Ts)) = T | |
| 554 |   fun find_double (t as Const (@{const_name fun_upd},T) $ f $ x $ y) =
 | |
| 555 | let | |
| 556 |       fun find (Const (@{const_name fun_upd},T) $ g $ v $ w) =
 | |
| 557 | if v aconv x then SOME g else gen_fun_upd (find g) T v w | |
| 558 | | find t = NONE | |
| 559 | in (dest_fun_T1 T, gen_fun_upd (find f) T x y) end | |
| 560 | fun fun_upd_prover ss = | |
| 561 | rtac eq_reflection 1 THEN rtac ext 1 THEN | |
| 562 |     simp_tac (Simplifier.inherit_context ss @{simpset}) 1
 | |
| 563 | val fun_upd2_simproc = | |
| 564 |     Simplifier.simproc @{theory}
 | |
| 565 | "fun_upd2" ["f(v := w, x := y)"] | |
| 566 | (fn _ => fn ss => fn t => | |
| 567 | case find_double t of (T, NONE) => NONE | |
| 568 | | (T, SOME rhs) => | |
| 569 | SOME (Goal.prove (Simplifier.the_context ss) [] [] | |
| 570 | (Term.equals T $ t $ rhs) (K (fun_upd_prover ss)))) | |
| 571 | in | |
| 572 | Addsimprocs [fun_upd2_simproc] | |
| 573 | end; | |
| 574 | *} | |
| 575 | ||
| 576 | ||
| 21870 | 577 | subsection {* Code generator setup *}
 | 
| 578 | ||
| 579 | code_const "op \<circ>" | |
| 580 | (SML infixl 5 "o") | |
| 581 | (Haskell infixr 9 ".") | |
| 582 | ||
| 21906 | 583 | code_const "id" | 
| 584 | (Haskell "id") | |
| 585 | ||
| 21870 | 586 | |
| 21547 
9c9fdf4c2949
moved order arities for fun and bool to Fun/Orderings
 haftmann parents: 
21327diff
changeset | 587 | subsection {* ML legacy bindings *} 
 | 
| 15510 | 588 | |
| 22845 | 589 | ML {*
 | 
| 590 | val set_cs = claset() delrules [equalityI] | |
| 591 | *} | |
| 5852 | 592 | |
| 22845 | 593 | ML {*
 | 
| 594 | val id_apply = @{thm id_apply}
 | |
| 595 | val id_def = @{thm id_def}
 | |
| 596 | val o_apply = @{thm o_apply}
 | |
| 597 | val o_assoc = @{thm o_assoc}
 | |
| 598 | val o_def = @{thm o_def}
 | |
| 599 | val injD = @{thm injD}
 | |
| 600 | val datatype_injI = @{thm datatype_injI}
 | |
| 601 | val range_ex1_eq = @{thm range_ex1_eq}
 | |
| 602 | val expand_fun_eq = @{thm expand_fun_eq}
 | |
| 13585 | 603 | *} | 
| 5852 | 604 | |
| 2912 | 605 | end |