| author | wenzelm | 
| Thu, 03 Jan 2019 20:16:42 +0100 | |
| changeset 69584 | a91e32843310 | 
| parent 69502 | 0cf906072e20 | 
| child 69593 | 3dda49e08b9d | 
| permissions | -rw-r--r-- | 
| 1475 | 1 | (* Title: HOL/Fun.thy | 
| 2 | Author: Tobias Nipkow, Cambridge University Computer Laboratory | |
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 3 | Author: Andrei Popescu, TU Muenchen | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 4 | Copyright 1994, 2012 | 
| 18154 | 5 | *) | 
| 923 | 6 | |
| 60758 | 7 | section \<open>Notions about functions\<close> | 
| 923 | 8 | |
| 15510 | 9 | theory Fun | 
| 63575 | 10 | imports Set | 
| 11 | keywords "functor" :: thy_goal | |
| 15131 | 12 | begin | 
| 2912 | 13 | |
| 63322 | 14 | lemma apply_inverse: "f x = u \<Longrightarrow> (\<And>x. P x \<Longrightarrow> g (f x) = x) \<Longrightarrow> P x \<Longrightarrow> x = g u" | 
| 26147 | 15 | by auto | 
| 2912 | 16 | |
| 63322 | 17 | text \<open>Uniqueness, so NOT the axiom of choice.\<close> | 
| 59504 
8c6747dba731
New lemmas and a bit of tidying up.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 18 | lemma uniq_choice: "\<forall>x. \<exists>!y. Q x y \<Longrightarrow> \<exists>f. \<forall>x. Q x (f x)" | 
| 
8c6747dba731
New lemmas and a bit of tidying up.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 19 | by (force intro: theI') | 
| 
8c6747dba731
New lemmas and a bit of tidying up.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 20 | |
| 
8c6747dba731
New lemmas and a bit of tidying up.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 21 | lemma b_uniq_choice: "\<forall>x\<in>S. \<exists>!y. Q x y \<Longrightarrow> \<exists>f. \<forall>x\<in>S. Q x (f x)" | 
| 
8c6747dba731
New lemmas and a bit of tidying up.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 22 | by (force intro: theI') | 
| 12258 | 23 | |
| 63400 | 24 | |
| 61799 | 25 | subsection \<open>The Identity Function \<open>id\<close>\<close> | 
| 6171 | 26 | |
| 63322 | 27 | definition id :: "'a \<Rightarrow> 'a" | 
| 28 | where "id = (\<lambda>x. x)" | |
| 13910 | 29 | |
| 26147 | 30 | lemma id_apply [simp]: "id x = x" | 
| 31 | by (simp add: id_def) | |
| 32 | ||
| 47579 | 33 | lemma image_id [simp]: "image id = id" | 
| 34 | by (simp add: id_def fun_eq_iff) | |
| 26147 | 35 | |
| 47579 | 36 | lemma vimage_id [simp]: "vimage id = id" | 
| 37 | by (simp add: id_def fun_eq_iff) | |
| 26147 | 38 | |
| 62843 
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 39 | lemma eq_id_iff: "(\<forall>x. f x = x) \<longleftrightarrow> f = id" | 
| 
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 40 | by auto | 
| 
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
 paulson <lp15@cam.ac.uk> parents: 
62618diff
changeset | 41 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51717diff
changeset | 42 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51717diff
changeset | 43 | constant id \<rightharpoonup> (Haskell) "id" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51717diff
changeset | 44 | |
| 26147 | 45 | |
| 61799 | 46 | subsection \<open>The Composition Operator \<open>f \<circ> g\<close>\<close> | 
| 26147 | 47 | |
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61799diff
changeset | 48 | definition comp :: "('b \<Rightarrow> 'c) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'c"  (infixl "\<circ>" 55)
 | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61799diff
changeset | 49 | where "f \<circ> g = (\<lambda>x. f (g x))" | 
| 11123 | 50 | |
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61799diff
changeset | 51 | notation (ASCII) | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61799diff
changeset | 52 | comp (infixl "o" 55) | 
| 19656 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19536diff
changeset | 53 | |
| 63322 | 54 | lemma comp_apply [simp]: "(f \<circ> g) x = f (g x)" | 
| 49739 | 55 | by (simp add: comp_def) | 
| 13585 | 56 | |
| 63322 | 57 | lemma comp_assoc: "(f \<circ> g) \<circ> h = f \<circ> (g \<circ> h)" | 
| 49739 | 58 | by (simp add: fun_eq_iff) | 
| 13585 | 59 | |
| 63322 | 60 | lemma id_comp [simp]: "id \<circ> g = g" | 
| 49739 | 61 | by (simp add: fun_eq_iff) | 
| 13585 | 62 | |
| 63322 | 63 | lemma comp_id [simp]: "f \<circ> id = f" | 
| 49739 | 64 | by (simp add: fun_eq_iff) | 
| 65 | ||
| 63575 | 66 | lemma comp_eq_dest: "a \<circ> b = c \<circ> d \<Longrightarrow> a (b v) = c (d v)" | 
| 49739 | 67 | by (simp add: fun_eq_iff) | 
| 34150 | 68 | |
| 63575 | 69 | lemma comp_eq_elim: "a \<circ> b = c \<circ> d \<Longrightarrow> ((\<And>v. a (b v) = c (d v)) \<Longrightarrow> R) \<Longrightarrow> R" | 
| 61204 | 70 | by (simp add: fun_eq_iff) | 
| 34150 | 71 | |
| 63322 | 72 | lemma comp_eq_dest_lhs: "a \<circ> b = c \<Longrightarrow> a (b v) = c v" | 
| 55066 | 73 | by clarsimp | 
| 74 | ||
| 63322 | 75 | lemma comp_eq_id_dest: "a \<circ> b = id \<circ> c \<Longrightarrow> a (b v) = c v" | 
| 76 | by clarsimp | |
| 77 | ||
| 78 | lemma image_comp: "f ` (g ` r) = (f \<circ> g) ` r" | |
| 33044 | 79 | by auto | 
| 80 | ||
| 63322 | 81 | lemma vimage_comp: "f -` (g -` x) = (g \<circ> f) -` x" | 
| 49739 | 82 | by auto | 
| 83 | ||
| 63322 | 84 | lemma image_eq_imp_comp: "f ` A = g ` B \<Longrightarrow> (h \<circ> f) ` A = (h \<circ> g) ` B" | 
| 59504 
8c6747dba731
New lemmas and a bit of tidying up.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 85 | by (auto simp: comp_def elim!: equalityE) | 
| 
8c6747dba731
New lemmas and a bit of tidying up.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 86 | |
| 67399 | 87 | lemma image_bind: "f ` (Set.bind A g) = Set.bind A ((`) f \<circ> g)" | 
| 63322 | 88 | by (auto simp add: Set.bind_def) | 
| 59512 | 89 | |
| 90 | lemma bind_image: "Set.bind (f ` A) g = Set.bind A (g \<circ> f)" | |
| 63322 | 91 | by (auto simp add: Set.bind_def) | 
| 59512 | 92 | |
| 63322 | 93 | lemma (in group_add) minus_comp_minus [simp]: "uminus \<circ> uminus = id" | 
| 60929 | 94 | by (simp add: fun_eq_iff) | 
| 95 | ||
| 63322 | 96 | lemma (in boolean_algebra) minus_comp_minus [simp]: "uminus \<circ> uminus = id" | 
| 60929 | 97 | by (simp add: fun_eq_iff) | 
| 98 | ||
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51717diff
changeset | 99 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51717diff
changeset | 100 | constant comp \<rightharpoonup> (SML) infixl 5 "o" and (Haskell) infixr 9 "." | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51717diff
changeset | 101 | |
| 13585 | 102 | |
| 61799 | 103 | subsection \<open>The Forward Composition Operator \<open>fcomp\<close>\<close> | 
| 26357 | 104 | |
| 63575 | 105 | definition fcomp :: "('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> 'a \<Rightarrow> 'c"  (infixl "\<circ>>" 60)
 | 
| 63322 | 106 | where "f \<circ>> g = (\<lambda>x. g (f x))" | 
| 26357 | 107 | |
| 37751 | 108 | lemma fcomp_apply [simp]: "(f \<circ>> g) x = g (f x)" | 
| 26357 | 109 | by (simp add: fcomp_def) | 
| 110 | ||
| 37751 | 111 | lemma fcomp_assoc: "(f \<circ>> g) \<circ>> h = f \<circ>> (g \<circ>> h)" | 
| 26357 | 112 | by (simp add: fcomp_def) | 
| 113 | ||
| 37751 | 114 | lemma id_fcomp [simp]: "id \<circ>> g = g" | 
| 26357 | 115 | by (simp add: fcomp_def) | 
| 116 | ||
| 37751 | 117 | lemma fcomp_id [simp]: "f \<circ>> id = f" | 
| 26357 | 118 | by (simp add: fcomp_def) | 
| 119 | ||
| 63322 | 120 | lemma fcomp_comp: "fcomp f g = comp g f" | 
| 61699 
a81dc5c4d6a9
New theorems mostly from Peter Gammie
 paulson <lp15@cam.ac.uk> parents: 
61630diff
changeset | 121 | by (simp add: ext) | 
| 
a81dc5c4d6a9
New theorems mostly from Peter Gammie
 paulson <lp15@cam.ac.uk> parents: 
61630diff
changeset | 122 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51717diff
changeset | 123 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51717diff
changeset | 124 | constant fcomp \<rightharpoonup> (Eval) infixl 1 "#>" | 
| 31202 
52d332f8f909
pretty printing of functional combinators for evaluation code
 haftmann parents: 
31080diff
changeset | 125 | |
| 37751 | 126 | no_notation fcomp (infixl "\<circ>>" 60) | 
| 26588 
d83271bfaba5
removed syntax from monad combinators; renamed mbind to scomp
 haftmann parents: 
26357diff
changeset | 127 | |
| 26357 | 128 | |
| 60758 | 129 | subsection \<open>Mapping functions\<close> | 
| 40602 | 130 | |
| 63322 | 131 | definition map_fun :: "('c \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'c \<Rightarrow> 'd"
 | 
| 132 | where "map_fun f g h = g \<circ> h \<circ> f" | |
| 40602 | 133 | |
| 63322 | 134 | lemma map_fun_apply [simp]: "map_fun f g h x = g (h (f x))" | 
| 40602 | 135 | by (simp add: map_fun_def) | 
| 136 | ||
| 137 | ||
| 60758 | 138 | subsection \<open>Injectivity and Bijectivity\<close> | 
| 39076 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 139 | |
| 63322 | 140 | definition inj_on :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> bool"  \<comment> \<open>injective\<close>
 | 
| 141 | where "inj_on f A \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>A. f x = f y \<longrightarrow> x = y)" | |
| 26147 | 142 | |
| 63322 | 143 | definition bij_betw :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> bool"  \<comment> \<open>bijective\<close>
 | 
| 144 | where "bij_betw f A B \<longleftrightarrow> inj_on f A \<and> f ` A = B" | |
| 26147 | 145 | |
| 63575 | 146 | text \<open> | 
| 147 | A common special case: functions injective, surjective or bijective over | |
| 148 | the entire domain type. | |
| 149 | \<close> | |
| 26147 | 150 | |
| 65170 
53675f36820d
restored surj as output abbreviation, amending 6af79184bef3
 haftmann parents: 
64966diff
changeset | 151 | abbreviation inj :: "('a \<Rightarrow> 'b) \<Rightarrow> bool"
 | 
| 
53675f36820d
restored surj as output abbreviation, amending 6af79184bef3
 haftmann parents: 
64966diff
changeset | 152 | where "inj f \<equiv> inj_on f UNIV" | 
| 26147 | 153 | |
| 65170 
53675f36820d
restored surj as output abbreviation, amending 6af79184bef3
 haftmann parents: 
64966diff
changeset | 154 | abbreviation surj :: "('a \<Rightarrow> 'b) \<Rightarrow> bool"
 | 
| 63322 | 155 | where "surj f \<equiv> range f = UNIV" | 
| 13585 | 156 | |
| 67226 | 157 | translations \<comment> \<open>The negated case:\<close> | 
| 65170 
53675f36820d
restored surj as output abbreviation, amending 6af79184bef3
 haftmann parents: 
64966diff
changeset | 158 | "\<not> CONST surj f" \<leftharpoondown> "CONST range f \<noteq> CONST UNIV" | 
| 
53675f36820d
restored surj as output abbreviation, amending 6af79184bef3
 haftmann parents: 
64966diff
changeset | 159 | |
| 
53675f36820d
restored surj as output abbreviation, amending 6af79184bef3
 haftmann parents: 
64966diff
changeset | 160 | abbreviation bij :: "('a \<Rightarrow> 'b) \<Rightarrow> bool"
 | 
| 
53675f36820d
restored surj as output abbreviation, amending 6af79184bef3
 haftmann parents: 
64966diff
changeset | 161 | where "bij f \<equiv> bij_betw f UNIV UNIV" | 
| 26147 | 162 | |
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 163 | lemma inj_def: "inj f \<longleftrightarrow> (\<forall>x y. f x = f y \<longrightarrow> x = y)" | 
| 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 164 | unfolding inj_on_def by blast | 
| 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 165 | |
| 63322 | 166 | lemma injI: "(\<And>x y. f x = f y \<Longrightarrow> x = y) \<Longrightarrow> inj f" | 
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 167 | unfolding inj_def by blast | 
| 13585 | 168 | |
| 63322 | 169 | theorem range_ex1_eq: "inj f \<Longrightarrow> b \<in> range f \<longleftrightarrow> (\<exists>!x. b = f x)" | 
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 170 | unfolding inj_def by blast | 
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 171 | |
| 63322 | 172 | lemma injD: "inj f \<Longrightarrow> f x = f y \<Longrightarrow> x = y" | 
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 173 | by (simp add: inj_def) | 
| 63322 | 174 | |
| 175 | lemma inj_on_eq_iff: "inj_on f A \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> f x = f y \<longleftrightarrow> x = y" | |
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 176 | by (auto simp: inj_on_def) | 
| 63322 | 177 | |
| 64965 | 178 | lemma inj_on_cong: "(\<And>a. a \<in> A \<Longrightarrow> f a = g a) \<Longrightarrow> inj_on f A \<longleftrightarrow> inj_on g A" | 
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 179 | by (auto simp: inj_on_def) | 
| 63322 | 180 | |
| 181 | lemma inj_on_strict_subset: "inj_on f B \<Longrightarrow> A \<subset> B \<Longrightarrow> f ` A \<subset> f ` B" | |
| 182 | unfolding inj_on_def by blast | |
| 183 | ||
| 184 | lemma inj_comp: "inj f \<Longrightarrow> inj g \<Longrightarrow> inj (f \<circ> g)" | |
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 185 | by (simp add: inj_def) | 
| 38620 | 186 | |
| 187 | lemma inj_fun: "inj f \<Longrightarrow> inj (\<lambda>x y. f x)" | |
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 188 | by (simp add: inj_def fun_eq_iff) | 
| 38620 | 189 | |
| 63322 | 190 | lemma inj_eq: "inj f \<Longrightarrow> f x = f y \<longleftrightarrow> x = y" | 
| 191 | by (simp add: inj_on_eq_iff) | |
| 32988 | 192 | |
| 26147 | 193 | lemma inj_on_id[simp]: "inj_on id A" | 
| 39076 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 194 | by (simp add: inj_on_def) | 
| 13585 | 195 | |
| 63322 | 196 | lemma inj_on_id2[simp]: "inj_on (\<lambda>x. x) A" | 
| 197 | by (simp add: inj_on_def) | |
| 26147 | 198 | |
| 46586 | 199 | lemma inj_on_Int: "inj_on f A \<or> inj_on f B \<Longrightarrow> inj_on f (A \<inter> B)" | 
| 63322 | 200 | unfolding inj_on_def by blast | 
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 201 | |
| 40702 | 202 | lemma surj_id: "surj id" | 
| 63322 | 203 | by simp | 
| 26147 | 204 | |
| 39101 
606432dd1896
Revert bij_betw changes to simp set (Problem in afp/Ordinals_and_Cardinals)
 hoelzl parents: 
39076diff
changeset | 205 | lemma bij_id[simp]: "bij id" | 
| 63322 | 206 | by (simp add: bij_betw_def) | 
| 13585 | 207 | |
| 63322 | 208 | lemma bij_uminus: "bij (uminus :: 'a \<Rightarrow> 'a::ab_group_add)" | 
| 209 | unfolding bij_betw_def inj_on_def | |
| 210 | by (force intro: minus_minus [symmetric]) | |
| 63072 | 211 | |
| 63322 | 212 | lemma inj_onI [intro?]: "(\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> f x = f y \<Longrightarrow> x = y) \<Longrightarrow> inj_on f A" | 
| 213 | by (simp add: inj_on_def) | |
| 13585 | 214 | |
| 63322 | 215 | lemma inj_on_inverseI: "(\<And>x. x \<in> A \<Longrightarrow> g (f x) = x) \<Longrightarrow> inj_on f A" | 
| 64965 | 216 | by (auto dest: arg_cong [of concl: g] simp add: inj_on_def) | 
| 13585 | 217 | |
| 63322 | 218 | lemma inj_onD: "inj_on f A \<Longrightarrow> f x = f y \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> x = y" | 
| 219 | unfolding inj_on_def by blast | |
| 13585 | 220 | |
| 63365 | 221 | lemma inj_on_subset: | 
| 222 | assumes "inj_on f A" | |
| 63575 | 223 | and "B \<subseteq> A" | 
| 63365 | 224 | shows "inj_on f B" | 
| 225 | proof (rule inj_onI) | |
| 226 | fix a b | |
| 227 | assume "a \<in> B" and "b \<in> B" | |
| 228 | with assms have "a \<in> A" and "b \<in> A" | |
| 229 | by auto | |
| 230 | moreover assume "f a = f b" | |
| 64965 | 231 | ultimately show "a = b" | 
| 232 | using assms by (auto dest: inj_onD) | |
| 63365 | 233 | qed | 
| 234 | ||
| 63322 | 235 | lemma comp_inj_on: "inj_on f A \<Longrightarrow> inj_on g (f ` A) \<Longrightarrow> inj_on (g \<circ> f) A" | 
| 236 | by (simp add: comp_def inj_on_def) | |
| 237 | ||
| 238 | lemma inj_on_imageI: "inj_on (g \<circ> f) A \<Longrightarrow> inj_on g (f ` A)" | |
| 63072 | 239 | by (auto simp add: inj_on_def) | 
| 15303 | 240 | |
| 63322 | 241 | lemma inj_on_image_iff: | 
| 64965 | 242 | "\<forall>x\<in>A. \<forall>y\<in>A. g (f x) = g (f y) \<longleftrightarrow> g x = g y \<Longrightarrow> inj_on f A \<Longrightarrow> inj_on g (f ` A) \<longleftrightarrow> inj_on g A" | 
| 63322 | 243 | unfolding inj_on_def by blast | 
| 15439 | 244 | |
| 63322 | 245 | lemma inj_on_contraD: "inj_on f A \<Longrightarrow> x \<noteq> y \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> f x \<noteq> f y" | 
| 246 | unfolding inj_on_def by blast | |
| 12258 | 247 | |
| 63072 | 248 | lemma inj_singleton [simp]: "inj_on (\<lambda>x. {x}) A"
 | 
| 249 | by (simp add: inj_on_def) | |
| 13585 | 250 | |
| 15111 | 251 | lemma inj_on_empty[iff]: "inj_on f {}"
 | 
| 63322 | 252 | by (simp add: inj_on_def) | 
| 13585 | 253 | |
| 63322 | 254 | lemma subset_inj_on: "inj_on f B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> inj_on f A" | 
| 255 | unfolding inj_on_def by blast | |
| 256 | ||
| 257 | lemma inj_on_Un: "inj_on f (A \<union> B) \<longleftrightarrow> inj_on f A \<and> inj_on f B \<and> f ` (A - B) \<inter> f ` (B - A) = {}"
 | |
| 258 | unfolding inj_on_def by (blast intro: sym) | |
| 15111 | 259 | |
| 63322 | 260 | lemma inj_on_insert [iff]: "inj_on f (insert a A) \<longleftrightarrow> inj_on f A \<and> f a \<notin> f ` (A - {a})"
 | 
| 261 | unfolding inj_on_def by (blast intro: sym) | |
| 262 | ||
| 263 | lemma inj_on_diff: "inj_on f A \<Longrightarrow> inj_on f (A - B)" | |
| 264 | unfolding inj_on_def by blast | |
| 15111 | 265 | |
| 63322 | 266 | lemma comp_inj_on_iff: "inj_on f A \<Longrightarrow> inj_on f' (f ` A) \<longleftrightarrow> inj_on (f' \<circ> f) A" | 
| 64965 | 267 | by (auto simp: comp_inj_on inj_on_def) | 
| 15111 | 268 | |
| 63322 | 269 | lemma inj_on_imageI2: "inj_on (f' \<circ> f) A \<Longrightarrow> inj_on f A" | 
| 64965 | 270 | by (auto simp: comp_inj_on inj_on_def) | 
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 271 | |
| 51598 
5dbe537087aa
generalized lemma fold_image thanks to Peter Lammich
 haftmann parents: 
49905diff
changeset | 272 | lemma inj_img_insertE: | 
| 
5dbe537087aa
generalized lemma fold_image thanks to Peter Lammich
 haftmann parents: 
49905diff
changeset | 273 | assumes "inj_on f A" | 
| 63322 | 274 | assumes "x \<notin> B" | 
| 275 | and "insert x B = f ` A" | |
| 276 | obtains x' A' where "x' \<notin> A'" and "A = insert x' A'" and "x = f x'" and "B = f ` A'" | |
| 51598 
5dbe537087aa
generalized lemma fold_image thanks to Peter Lammich
 haftmann parents: 
49905diff
changeset | 277 | proof - | 
| 
5dbe537087aa
generalized lemma fold_image thanks to Peter Lammich
 haftmann parents: 
49905diff
changeset | 278 | from assms have "x \<in> f ` A" by auto | 
| 
5dbe537087aa
generalized lemma fold_image thanks to Peter Lammich
 haftmann parents: 
49905diff
changeset | 279 | then obtain x' where *: "x' \<in> A" "x = f x'" by auto | 
| 63322 | 280 |   then have A: "A = insert x' (A - {x'})" by auto
 | 
| 281 |   with assms * have B: "B = f ` (A - {x'})" by (auto dest: inj_on_contraD)
 | |
| 51598 
5dbe537087aa
generalized lemma fold_image thanks to Peter Lammich
 haftmann parents: 
49905diff
changeset | 282 |   have "x' \<notin> A - {x'}" by simp
 | 
| 63322 | 283 | from this A \<open>x = f x'\<close> B show ?thesis .. | 
| 51598 
5dbe537087aa
generalized lemma fold_image thanks to Peter Lammich
 haftmann parents: 
49905diff
changeset | 284 | qed | 
| 
5dbe537087aa
generalized lemma fold_image thanks to Peter Lammich
 haftmann parents: 
49905diff
changeset | 285 | |
| 54578 
9387251b6a46
eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
 traytel parents: 
54147diff
changeset | 286 | lemma linorder_injI: | 
| 64965 | 287 | assumes "\<And>x y::'a::linorder. x < y \<Longrightarrow> f x \<noteq> f y" | 
| 54578 
9387251b6a46
eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
 traytel parents: 
54147diff
changeset | 288 | shows "inj f" | 
| 61799 | 289 | \<comment> \<open>Courtesy of Stephan Merz\<close> | 
| 54578 
9387251b6a46
eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
 traytel parents: 
54147diff
changeset | 290 | proof (rule inj_onI) | 
| 63400 | 291 | show "x = y" if "f x = f y" for x y | 
| 64965 | 292 | by (rule linorder_cases) (auto dest: assms simp: that) | 
| 54578 
9387251b6a46
eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
 traytel parents: 
54147diff
changeset | 293 | qed | 
| 
9387251b6a46
eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
 traytel parents: 
54147diff
changeset | 294 | |
| 40702 | 295 | lemma surj_def: "surj f \<longleftrightarrow> (\<forall>y. \<exists>x. y = f x)" | 
| 296 | by auto | |
| 39076 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 297 | |
| 63322 | 298 | lemma surjI: | 
| 64965 | 299 | assumes "\<And>x. g (f x) = x" | 
| 63322 | 300 | shows "surj g" | 
| 64965 | 301 | using assms [symmetric] by auto | 
| 13585 | 302 | |
| 39076 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 303 | lemma surjD: "surj f \<Longrightarrow> \<exists>x. y = f x" | 
| 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 304 | by (simp add: surj_def) | 
| 13585 | 305 | |
| 39076 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 306 | lemma surjE: "surj f \<Longrightarrow> (\<And>x. y = f x \<Longrightarrow> C) \<Longrightarrow> C" | 
| 63575 | 307 | by (simp add: surj_def) blast | 
| 13585 | 308 | |
| 63322 | 309 | lemma comp_surj: "surj f \<Longrightarrow> surj g \<Longrightarrow> surj (g \<circ> f)" | 
| 63416 
6af79184bef3
avoid to hide equality behind (output) abbreviation
 haftmann parents: 
63400diff
changeset | 310 | by (simp add: image_comp [symmetric]) | 
| 13585 | 311 | |
| 63322 | 312 | lemma bij_betw_imageI: "inj_on f A \<Longrightarrow> f ` A = B \<Longrightarrow> bij_betw f A B" | 
| 313 | unfolding bij_betw_def by clarify | |
| 57282 | 314 | |
| 315 | lemma bij_betw_imp_surj_on: "bij_betw f A B \<Longrightarrow> f ` A = B" | |
| 316 | unfolding bij_betw_def by clarify | |
| 317 | ||
| 39074 | 318 | lemma bij_betw_imp_surj: "bij_betw f A UNIV \<Longrightarrow> surj f" | 
| 40702 | 319 | unfolding bij_betw_def by auto | 
| 39074 | 320 | |
| 63322 | 321 | lemma bij_betw_empty1: "bij_betw f {} A \<Longrightarrow> A = {}"
 | 
| 322 | unfolding bij_betw_def by blast | |
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 323 | |
| 63322 | 324 | lemma bij_betw_empty2: "bij_betw f A {} \<Longrightarrow> A = {}"
 | 
| 325 | unfolding bij_betw_def by blast | |
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 326 | |
| 63322 | 327 | lemma inj_on_imp_bij_betw: "inj_on f A \<Longrightarrow> bij_betw f A (f ` A)" | 
| 328 | unfolding bij_betw_def by simp | |
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 329 | |
| 39076 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 330 | lemma bij_def: "bij f \<longleftrightarrow> inj f \<and> surj f" | 
| 64965 | 331 | by (rule bij_betw_def) | 
| 39074 | 332 | |
| 63322 | 333 | lemma bijI: "inj f \<Longrightarrow> surj f \<Longrightarrow> bij f" | 
| 64965 | 334 | by (rule bij_betw_imageI) | 
| 13585 | 335 | |
| 63322 | 336 | lemma bij_is_inj: "bij f \<Longrightarrow> inj f" | 
| 337 | by (simp add: bij_def) | |
| 13585 | 338 | |
| 63322 | 339 | lemma bij_is_surj: "bij f \<Longrightarrow> surj f" | 
| 340 | by (simp add: bij_def) | |
| 13585 | 341 | |
| 26105 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 342 | lemma bij_betw_imp_inj_on: "bij_betw f A B \<Longrightarrow> inj_on f A" | 
| 63322 | 343 | by (simp add: bij_betw_def) | 
| 26105 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 344 | |
| 63322 | 345 | lemma bij_betw_trans: "bij_betw f A B \<Longrightarrow> bij_betw g B C \<Longrightarrow> bij_betw (g \<circ> f) A C" | 
| 346 | by (auto simp add:bij_betw_def comp_inj_on) | |
| 31438 | 347 | |
| 63322 | 348 | lemma bij_comp: "bij f \<Longrightarrow> bij g \<Longrightarrow> bij (g \<circ> f)" | 
| 40702 | 349 | by (rule bij_betw_trans) | 
| 350 | ||
| 63322 | 351 | lemma bij_betw_comp_iff: "bij_betw f A A' \<Longrightarrow> bij_betw f' A' A'' \<longleftrightarrow> bij_betw (f' \<circ> f) A A''" | 
| 352 | by (auto simp add: bij_betw_def inj_on_def) | |
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 353 | |
| 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 354 | lemma bij_betw_comp_iff2: | 
| 63322 | 355 | assumes bij: "bij_betw f' A' A''" | 
| 356 | and img: "f ` A \<le> A'" | |
| 357 | shows "bij_betw f A A' \<longleftrightarrow> bij_betw (f' \<circ> f) A A''" | |
| 358 | using assms | |
| 359 | proof (auto simp add: bij_betw_comp_iff) | |
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 360 | assume *: "bij_betw (f' \<circ> f) A A''" | 
| 63322 | 361 | then show "bij_betw f A A'" | 
| 362 | using img | |
| 363 | proof (auto simp add: bij_betw_def) | |
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 364 | assume "inj_on (f' \<circ> f) A" | 
| 63575 | 365 | then show "inj_on f A" | 
| 366 | using inj_on_imageI2 by blast | |
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 367 | next | 
| 63322 | 368 | fix a' | 
| 369 | assume **: "a' \<in> A'" | |
| 63575 | 370 | with bij have "f' a' \<in> A''" | 
| 371 | unfolding bij_betw_def by auto | |
| 372 | with * obtain a where 1: "a \<in> A \<and> f' (f a) = f' a'" | |
| 373 | unfolding bij_betw_def by force | |
| 374 | with img have "f a \<in> A'" by auto | |
| 375 | with bij ** 1 have "f a = a'" | |
| 376 | unfolding bij_betw_def inj_on_def by auto | |
| 377 | with 1 show "a' \<in> f ` A" by auto | |
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 378 | qed | 
| 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 379 | qed | 
| 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 380 | |
| 63322 | 381 | lemma bij_betw_inv: | 
| 382 | assumes "bij_betw f A B" | |
| 383 | shows "\<exists>g. bij_betw g B A" | |
| 26105 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 384 | proof - | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 385 | have i: "inj_on f A" and s: "f ` A = B" | 
| 63322 | 386 | using assms by (auto simp: bij_betw_def) | 
| 387 | let ?P = "\<lambda>b a. a \<in> A \<and> f a = b" | |
| 388 | let ?g = "\<lambda>b. The (?P b)" | |
| 389 | have g: "?g b = a" if P: "?P b a" for a b | |
| 390 | proof - | |
| 63575 | 391 | from that s have ex1: "\<exists>a. ?P b a" by blast | 
| 63322 | 392 | then have uex1: "\<exists>!a. ?P b a" by (blast dest:inj_onD[OF i]) | 
| 63575 | 393 | then show ?thesis | 
| 394 | using the1_equality[OF uex1, OF P] P by simp | |
| 63322 | 395 | qed | 
| 26105 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 396 | have "inj_on ?g B" | 
| 63322 | 397 | proof (rule inj_onI) | 
| 398 | fix x y | |
| 399 | assume "x \<in> B" "y \<in> B" "?g x = ?g y" | |
| 400 | from s \<open>x \<in> B\<close> obtain a1 where a1: "?P x a1" by blast | |
| 401 | from s \<open>y \<in> B\<close> obtain a2 where a2: "?P y a2" by blast | |
| 402 | from g [OF a1] a1 g [OF a2] a2 \<open>?g x = ?g y\<close> show "x = y" by simp | |
| 26105 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 403 | qed | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 404 | moreover have "?g ` B = A" | 
| 63322 | 405 | proof (auto simp: image_def) | 
| 406 | fix b | |
| 407 | assume "b \<in> B" | |
| 56077 | 408 | with s obtain a where P: "?P b a" by blast | 
| 63575 | 409 | with g[OF P] show "?g b \<in> A" by auto | 
| 26105 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 410 | next | 
| 63322 | 411 | fix a | 
| 412 | assume "a \<in> A" | |
| 63575 | 413 | with s obtain b where P: "?P b a" by blast | 
| 414 | with s have "b \<in> B" by blast | |
| 26105 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 415 | with g[OF P] show "\<exists>b\<in>B. a = ?g b" by blast | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 416 | qed | 
| 63575 | 417 | ultimately show ?thesis | 
| 418 | by (auto simp: bij_betw_def) | |
| 26105 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 419 | qed | 
| 
ae06618225ec
moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
 nipkow parents: 
25886diff
changeset | 420 | |
| 63588 | 421 | lemma bij_betw_cong: "(\<And>a. a \<in> A \<Longrightarrow> f a = g a) \<Longrightarrow> bij_betw f A A' = bij_betw g A A'" | 
| 63591 | 422 | unfolding bij_betw_def inj_on_def by safe force+ (* somewhat slow *) | 
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 423 | |
| 63322 | 424 | lemma bij_betw_id[intro, simp]: "bij_betw id A A" | 
| 425 | unfolding bij_betw_def id_def by auto | |
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 426 | |
| 63322 | 427 | lemma bij_betw_id_iff: "bij_betw id A B \<longleftrightarrow> A = B" | 
| 428 | by (auto simp add: bij_betw_def) | |
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 429 | |
| 39075 | 430 | lemma bij_betw_combine: | 
| 63400 | 431 |   "bij_betw f A B \<Longrightarrow> bij_betw f C D \<Longrightarrow> B \<inter> D = {} \<Longrightarrow> bij_betw f (A \<union> C) (B \<union> D)"
 | 
| 432 | unfolding bij_betw_def inj_on_Un image_Un by auto | |
| 39075 | 433 | |
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 434 | lemma bij_betw_subset: "bij_betw f A A' \<Longrightarrow> B \<subseteq> A \<Longrightarrow> f ` B = B' \<Longrightarrow> bij_betw f B B'" | 
| 63322 | 435 | by (auto simp add: bij_betw_def inj_on_def) | 
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 436 | |
| 58195 | 437 | lemma bij_pointE: | 
| 438 | assumes "bij f" | |
| 439 | obtains x where "y = f x" and "\<And>x'. y = f x' \<Longrightarrow> x' = x" | |
| 440 | proof - | |
| 441 | from assms have "inj f" by (rule bij_is_inj) | |
| 442 | moreover from assms have "surj f" by (rule bij_is_surj) | |
| 443 | then have "y \<in> range f" by simp | |
| 444 | ultimately have "\<exists>!x. y = f x" by (simp add: range_ex1_eq) | |
| 445 | with that show thesis by blast | |
| 446 | qed | |
| 447 | ||
| 63322 | 448 | lemma surj_image_vimage_eq: "surj f \<Longrightarrow> f ` (f -` A) = A" | 
| 449 | by simp | |
| 13585 | 450 | |
| 42903 | 451 | lemma surj_vimage_empty: | 
| 63322 | 452 | assumes "surj f" | 
| 453 |   shows "f -` A = {} \<longleftrightarrow> A = {}"
 | |
| 454 | using surj_image_vimage_eq [OF \<open>surj f\<close>, of A] | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44860diff
changeset | 455 | by (intro iffI) fastforce+ | 
| 42903 | 456 | |
| 63322 | 457 | lemma inj_vimage_image_eq: "inj f \<Longrightarrow> f -` (f ` A) = A" | 
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 458 | unfolding inj_def by blast | 
| 13585 | 459 | |
| 63322 | 460 | lemma vimage_subsetD: "surj f \<Longrightarrow> f -` B \<subseteq> A \<Longrightarrow> B \<subseteq> f ` A" | 
| 461 | by (blast intro: sym) | |
| 13585 | 462 | |
| 63322 | 463 | lemma vimage_subsetI: "inj f \<Longrightarrow> B \<subseteq> f ` A \<Longrightarrow> f -` B \<subseteq> A" | 
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 464 | unfolding inj_def by blast | 
| 13585 | 465 | |
| 63322 | 466 | lemma vimage_subset_eq: "bij f \<Longrightarrow> f -` B \<subseteq> A \<longleftrightarrow> B \<subseteq> f ` A" | 
| 467 | unfolding bij_def by (blast del: subsetI intro: vimage_subsetI vimage_subsetD) | |
| 13585 | 468 | |
| 63322 | 469 | lemma inj_on_image_eq_iff: "inj_on f C \<Longrightarrow> A \<subseteq> C \<Longrightarrow> B \<subseteq> C \<Longrightarrow> f ` A = f ` B \<longleftrightarrow> A = B" | 
| 64965 | 470 | by (fastforce simp: inj_on_def) | 
| 53927 | 471 | |
| 31438 | 472 | lemma inj_on_Un_image_eq_iff: "inj_on f (A \<union> B) \<Longrightarrow> f ` A = f ` B \<longleftrightarrow> A = B" | 
| 63322 | 473 | by (erule inj_on_image_eq_iff) simp_all | 
| 31438 | 474 | |
| 63322 | 475 | lemma inj_on_image_Int: "inj_on f C \<Longrightarrow> A \<subseteq> C \<Longrightarrow> B \<subseteq> C \<Longrightarrow> f ` (A \<inter> B) = f ` A \<inter> f ` B" | 
| 476 | unfolding inj_on_def by blast | |
| 477 | ||
| 478 | lemma inj_on_image_set_diff: "inj_on f C \<Longrightarrow> A - B \<subseteq> C \<Longrightarrow> B \<subseteq> C \<Longrightarrow> f ` (A - B) = f ` A - f ` B" | |
| 479 | unfolding inj_on_def by blast | |
| 13585 | 480 | |
| 63322 | 481 | lemma image_Int: "inj f \<Longrightarrow> f ` (A \<inter> B) = f ` A \<inter> f ` B" | 
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 482 | unfolding inj_def by blast | 
| 13585 | 483 | |
| 63322 | 484 | lemma image_set_diff: "inj f \<Longrightarrow> f ` (A - B) = f ` A - f ` B" | 
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 485 | unfolding inj_def by blast | 
| 13585 | 486 | |
| 63322 | 487 | lemma inj_on_image_mem_iff: "inj_on f B \<Longrightarrow> a \<in> B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> f a \<in> f ` A \<longleftrightarrow> a \<in> A" | 
| 59504 
8c6747dba731
New lemmas and a bit of tidying up.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 488 | by (auto simp: inj_on_def) | 
| 
8c6747dba731
New lemmas and a bit of tidying up.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 489 | |
| 61520 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61378diff
changeset | 490 | (*FIXME DELETE*) | 
| 63322 | 491 | lemma inj_on_image_mem_iff_alt: "inj_on f B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> f a \<in> f ` A \<Longrightarrow> a \<in> B \<Longrightarrow> a \<in> A" | 
| 61520 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61378diff
changeset | 492 | by (blast dest: inj_onD) | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61378diff
changeset | 493 | |
| 63322 | 494 | lemma inj_image_mem_iff: "inj f \<Longrightarrow> f a \<in> f ` A \<longleftrightarrow> a \<in> A" | 
| 59504 
8c6747dba731
New lemmas and a bit of tidying up.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 495 | by (blast dest: injD) | 
| 13585 | 496 | |
| 63322 | 497 | lemma inj_image_subset_iff: "inj f \<Longrightarrow> f ` A \<subseteq> f ` B \<longleftrightarrow> A \<subseteq> B" | 
| 59504 
8c6747dba731
New lemmas and a bit of tidying up.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 498 | by (blast dest: injD) | 
| 13585 | 499 | |
| 63322 | 500 | lemma inj_image_eq_iff: "inj f \<Longrightarrow> f ` A = f ` B \<longleftrightarrow> A = B" | 
| 59504 
8c6747dba731
New lemmas and a bit of tidying up.
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 501 | by (blast dest: injD) | 
| 13585 | 502 | |
| 63322 | 503 | lemma surj_Compl_image_subset: "surj f \<Longrightarrow> - (f ` A) \<subseteq> f ` (- A)" | 
| 504 | by auto | |
| 5852 | 505 | |
| 63322 | 506 | lemma inj_image_Compl_subset: "inj f \<Longrightarrow> f ` (- A) \<subseteq> - (f ` A)" | 
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 507 | by (auto simp: inj_def) | 
| 63322 | 508 | |
| 509 | lemma bij_image_Compl_eq: "bij f \<Longrightarrow> f ` (- A) = - (f ` A)" | |
| 510 | by (simp add: bij_def inj_image_Compl_subset surj_Compl_image_subset equalityI) | |
| 13585 | 511 | |
| 41657 | 512 | lemma inj_vimage_singleton: "inj f \<Longrightarrow> f -` {a} \<subseteq> {THE x. f x = a}"
 | 
| 63322 | 513 | \<comment> \<open>The inverse image of a singleton under an injective function is included in a singleton.\<close> | 
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 514 | by (simp add: inj_def) (blast intro: the_equality [symmetric]) | 
| 41657 | 515 | |
| 63322 | 516 | lemma inj_on_vimage_singleton: "inj_on f A \<Longrightarrow> f -` {a} \<inter> A \<subseteq> {THE x. x \<in> A \<and> f x = a}"
 | 
| 43991 | 517 | by (auto simp add: inj_on_def intro: the_equality [symmetric]) | 
| 518 | ||
| 35584 
768f8d92b767
generalized inj_uminus; added strict_mono_imp_inj_on
 hoelzl parents: 
35580diff
changeset | 519 | lemma (in ordered_ab_group_add) inj_uminus[simp, intro]: "inj_on uminus A" | 
| 35580 | 520 | by (auto intro!: inj_onI) | 
| 13585 | 521 | |
| 35584 
768f8d92b767
generalized inj_uminus; added strict_mono_imp_inj_on
 hoelzl parents: 
35580diff
changeset | 522 | lemma (in linorder) strict_mono_imp_inj_on: "strict_mono f \<Longrightarrow> inj_on f A" | 
| 
768f8d92b767
generalized inj_uminus; added strict_mono_imp_inj_on
 hoelzl parents: 
35580diff
changeset | 523 | by (auto intro!: inj_onI dest: strict_mono_eq) | 
| 
768f8d92b767
generalized inj_uminus; added strict_mono_imp_inj_on
 hoelzl parents: 
35580diff
changeset | 524 | |
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 525 | lemma bij_betw_byWitness: | 
| 63322 | 526 | assumes left: "\<forall>a \<in> A. f' (f a) = a" | 
| 527 | and right: "\<forall>a' \<in> A'. f (f' a') = a'" | |
| 63575 | 528 | and "f ` A \<subseteq> A'" | 
| 529 | and img2: "f' ` A' \<subseteq> A" | |
| 63322 | 530 | shows "bij_betw f A A'" | 
| 531 | using assms | |
| 63400 | 532 | unfolding bij_betw_def inj_on_def | 
| 533 | proof safe | |
| 63322 | 534 | fix a b | 
| 63575 | 535 | assume "a \<in> A" "b \<in> A" | 
| 536 | with left have "a = f' (f a) \<and> b = f' (f b)" by simp | |
| 537 | moreover assume "f a = f b" | |
| 538 | ultimately show "a = b" by simp | |
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 539 | next | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 540 | fix a' assume *: "a' \<in> A'" | 
| 63575 | 541 | with img2 have "f' a' \<in> A" by blast | 
| 542 | moreover from * right have "a' = f (f' a')" by simp | |
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 543 | ultimately show "a' \<in> f ` A" by blast | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 544 | qed | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 545 | |
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 546 | corollary notIn_Un_bij_betw: | 
| 63322 | 547 | assumes "b \<notin> A" | 
| 548 | and "f b \<notin> A'" | |
| 549 | and "bij_betw f A A'" | |
| 550 |   shows "bij_betw f (A \<union> {b}) (A' \<union> {f b})"
 | |
| 551 | proof - | |
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 552 |   have "bij_betw f {b} {f b}"
 | 
| 63322 | 553 | unfolding bij_betw_def inj_on_def by simp | 
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 554 | with assms show ?thesis | 
| 63322 | 555 |     using bij_betw_combine[of f A A' "{b}" "{f b}"] by blast
 | 
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 556 | qed | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 557 | |
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 558 | lemma notIn_Un_bij_betw3: | 
| 63322 | 559 | assumes "b \<notin> A" | 
| 560 | and "f b \<notin> A'" | |
| 561 |   shows "bij_betw f A A' = bij_betw f (A \<union> {b}) (A' \<union> {f b})"
 | |
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 562 | proof | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 563 | assume "bij_betw f A A'" | 
| 63322 | 564 |   then show "bij_betw f (A \<union> {b}) (A' \<union> {f b})"
 | 
| 565 | using assms notIn_Un_bij_betw [of b A f A'] by blast | |
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 566 | next | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 567 |   assume *: "bij_betw f (A \<union> {b}) (A' \<union> {f b})"
 | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 568 | have "f ` A = A'" | 
| 63322 | 569 | proof auto | 
| 570 | fix a | |
| 571 | assume **: "a \<in> A" | |
| 572 |     then have "f a \<in> A' \<union> {f b}"
 | |
| 573 | using * unfolding bij_betw_def by blast | |
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 574 | moreover | 
| 63322 | 575 | have False if "f a = f b" | 
| 576 | proof - | |
| 63575 | 577 | have "a = b" | 
| 578 | using * ** that unfolding bij_betw_def inj_on_def by blast | |
| 63322 | 579 | with \<open>b \<notin> A\<close> ** show ?thesis by blast | 
| 580 | qed | |
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 581 | ultimately show "f a \<in> A'" by blast | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 582 | next | 
| 63322 | 583 | fix a' | 
| 584 | assume **: "a' \<in> A'" | |
| 585 |     then have "a' \<in> f ` (A \<union> {b})"
 | |
| 586 | using * by (auto simp add: bij_betw_def) | |
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 587 |     then obtain a where 1: "a \<in> A \<union> {b} \<and> f a = a'" by blast
 | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 588 | moreover | 
| 63322 | 589 | have False if "a = b" using 1 ** \<open>f b \<notin> A'\<close> that by blast | 
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 590 | ultimately have "a \<in> A" by blast | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 591 | with 1 show "a' \<in> f ` A" by blast | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 592 | qed | 
| 63322 | 593 | then show "bij_betw f A A'" | 
| 594 |     using * bij_betw_subset[of f "A \<union> {b}" _ A] by blast
 | |
| 55019 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 595 | qed | 
| 
0d5e831175de
moved lemmas from 'Fun_More_FP' to where they belong
 blanchet parents: 
54578diff
changeset | 596 | |
| 69502 | 597 | text \<open>Important examples\<close> | 
| 598 | ||
| 599 | context cancel_semigroup_add | |
| 600 | begin | |
| 601 | ||
| 602 | lemma inj_add_left [simp]: \<open>inj ((+) a)\<close> | |
| 603 | by (rule injI) simp | |
| 604 | ||
| 605 | end | |
| 606 | ||
| 607 | context ab_group_add | |
| 608 | begin | |
| 609 | ||
| 610 | lemma inj_diff_right [simp]: \<open>inj (\<lambda>b. b - a)\<close> | |
| 611 | proof - | |
| 612 | have \<open>inj ((+) (- a))\<close> | |
| 613 | by (fact inj_add_left) | |
| 614 | also have \<open>(+) (- a) = (\<lambda>b. b - a)\<close> | |
| 615 | by (simp add: fun_eq_iff) | |
| 616 | finally show ?thesis . | |
| 617 | qed | |
| 618 | ||
| 619 | end | |
| 620 | ||
| 41657 | 621 | |
| 63322 | 622 | subsection \<open>Function Updating\<close> | 
| 13585 | 623 | |
| 63322 | 624 | definition fun_upd :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a \<Rightarrow> 'b)"
 | 
| 63324 | 625 | where "fun_upd f a b = (\<lambda>x. if x = a then b else f x)" | 
| 26147 | 626 | |
| 41229 
d797baa3d57c
replaced command 'nonterminals' by slightly modernized version 'nonterminal';
 wenzelm parents: 
40969diff
changeset | 627 | nonterminal updbinds and updbind | 
| 
d797baa3d57c
replaced command 'nonterminals' by slightly modernized version 'nonterminal';
 wenzelm parents: 
40969diff
changeset | 628 | |
| 26147 | 629 | syntax | 
| 63322 | 630 |   "_updbind" :: "'a \<Rightarrow> 'a \<Rightarrow> updbind"             ("(2_ :=/ _)")
 | 
| 631 |   ""         :: "updbind \<Rightarrow> updbinds"             ("_")
 | |
| 632 |   "_updbinds":: "updbind \<Rightarrow> updbinds \<Rightarrow> updbinds" ("_,/ _")
 | |
| 633 |   "_Update"  :: "'a \<Rightarrow> updbinds \<Rightarrow> 'a"            ("_/'((_)')" [1000, 0] 900)
 | |
| 26147 | 634 | |
| 635 | translations | |
| 63322 | 636 | "_Update f (_updbinds b bs)" \<rightleftharpoons> "_Update (_Update f b) bs" | 
| 637 | "f(x:=y)" \<rightleftharpoons> "CONST fun_upd f x y" | |
| 26147 | 638 | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55066diff
changeset | 639 | (* Hint: to define the sum of two functions (or maps), use case_sum. | 
| 58111 | 640 | A nice infix syntax could be defined by | 
| 35115 | 641 | notation | 
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55066diff
changeset | 642 | case_sum (infixr "'(+')"80) | 
| 26147 | 643 | *) | 
| 644 | ||
| 63322 | 645 | lemma fun_upd_idem_iff: "f(x:=y) = f \<longleftrightarrow> f x = y" | 
| 646 | unfolding fun_upd_def | |
| 647 | apply safe | |
| 63575 | 648 | apply (erule subst) | 
| 649 | apply (rule_tac [2] ext) | |
| 650 | apply auto | |
| 63322 | 651 | done | 
| 13585 | 652 | |
| 63322 | 653 | lemma fun_upd_idem: "f x = y \<Longrightarrow> f(x := y) = f" | 
| 45603 | 654 | by (simp only: fun_upd_idem_iff) | 
| 13585 | 655 | |
| 45603 | 656 | lemma fun_upd_triv [iff]: "f(x := f x) = f" | 
| 657 | by (simp only: fun_upd_idem) | |
| 13585 | 658 | |
| 63322 | 659 | lemma fun_upd_apply [simp]: "(f(x := y)) z = (if z = x then y else f z)" | 
| 660 | by (simp add: fun_upd_def) | |
| 13585 | 661 | |
| 63322 | 662 | (* fun_upd_apply supersedes these two, but they are useful | 
| 13585 | 663 | if fun_upd_apply is intentionally removed from the simpset *) | 
| 63322 | 664 | lemma fun_upd_same: "(f(x := y)) x = y" | 
| 665 | by simp | |
| 13585 | 666 | |
| 63322 | 667 | lemma fun_upd_other: "z \<noteq> x \<Longrightarrow> (f(x := y)) z = f z" | 
| 668 | by simp | |
| 13585 | 669 | |
| 63322 | 670 | lemma fun_upd_upd [simp]: "f(x := y, x := z) = f(x := z)" | 
| 671 | by (simp add: fun_eq_iff) | |
| 13585 | 672 | |
| 63322 | 673 | lemma fun_upd_twist: "a \<noteq> c \<Longrightarrow> (m(a := b))(c := d) = (m(c := d))(a := b)" | 
| 674 | by (rule ext) auto | |
| 675 | ||
| 676 | lemma inj_on_fun_updI: "inj_on f A \<Longrightarrow> y \<notin> f ` A \<Longrightarrow> inj_on (f(x := y)) A" | |
| 64966 
d53d7ca3303e
added inj_def (redundant, analogous to surj_def, bij_def);
 wenzelm parents: 
64965diff
changeset | 677 | by (auto simp: inj_on_def) | 
| 15303 | 678 | |
| 63322 | 679 | lemma fun_upd_image: "f(x := y) ` A = (if x \<in> A then insert y (f ` (A - {x})) else f ` A)"
 | 
| 680 | by auto | |
| 15510 | 681 | |
| 31080 | 682 | lemma fun_upd_comp: "f \<circ> (g(x := y)) = (f \<circ> g)(x := f y)" | 
| 44921 | 683 | by auto | 
| 31080 | 684 | |
| 61630 | 685 | lemma fun_upd_eqD: "f(x := y) = g(x := z) \<Longrightarrow> y = z" | 
| 63322 | 686 | by (simp add: fun_eq_iff split: if_split_asm) | 
| 687 | ||
| 26147 | 688 | |
| 61799 | 689 | subsection \<open>\<open>override_on\<close>\<close> | 
| 26147 | 690 | |
| 63322 | 691 | definition override_on :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> 'b"
 | 
| 692 | where "override_on f g A = (\<lambda>a. if a \<in> A then g a else f a)" | |
| 13910 | 693 | |
| 15691 | 694 | lemma override_on_emptyset[simp]: "override_on f g {} = f"
 | 
| 64965 | 695 | by (simp add: override_on_def) | 
| 13910 | 696 | |
| 63322 | 697 | lemma override_on_apply_notin[simp]: "a \<notin> A \<Longrightarrow> (override_on f g A) a = f a" | 
| 64965 | 698 | by (simp add: override_on_def) | 
| 13910 | 699 | |
| 63322 | 700 | lemma override_on_apply_in[simp]: "a \<in> A \<Longrightarrow> (override_on f g A) a = g a" | 
| 64965 | 701 | by (simp add: override_on_def) | 
| 13910 | 702 | |
| 63561 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63416diff
changeset | 703 | lemma override_on_insert: "override_on f g (insert x X) = (override_on f g X)(x:=g x)" | 
| 64965 | 704 | by (simp add: override_on_def fun_eq_iff) | 
| 63561 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63416diff
changeset | 705 | |
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63416diff
changeset | 706 | lemma override_on_insert': "override_on f g (insert x X) = (override_on (f(x:=g x)) g X)" | 
| 64965 | 707 | by (simp add: override_on_def fun_eq_iff) | 
| 63561 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
63416diff
changeset | 708 | |
| 26147 | 709 | |
| 61799 | 710 | subsection \<open>\<open>swap\<close>\<close> | 
| 15510 | 711 | |
| 56608 | 712 | definition swap :: "'a \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b)"
 | 
| 63322 | 713 | where "swap a b f = f (a := f b, b:= f a)" | 
| 15510 | 714 | |
| 56608 | 715 | lemma swap_apply [simp]: | 
| 716 | "swap a b f a = f b" | |
| 717 | "swap a b f b = f a" | |
| 718 | "c \<noteq> a \<Longrightarrow> c \<noteq> b \<Longrightarrow> swap a b f c = f c" | |
| 719 | by (simp_all add: swap_def) | |
| 720 | ||
| 63322 | 721 | lemma swap_self [simp]: "swap a a f = f" | 
| 56608 | 722 | by (simp add: swap_def) | 
| 15510 | 723 | |
| 63322 | 724 | lemma swap_commute: "swap a b f = swap b a f" | 
| 56608 | 725 | by (simp add: fun_upd_def swap_def fun_eq_iff) | 
| 15510 | 726 | |
| 63322 | 727 | lemma swap_nilpotent [simp]: "swap a b (swap a b f) = f" | 
| 63575 | 728 | by (rule ext) (simp add: fun_upd_def swap_def) | 
| 56608 | 729 | |
| 63322 | 730 | lemma swap_comp_involutory [simp]: "swap a b \<circ> swap a b = id" | 
| 56608 | 731 | by (rule ext) simp | 
| 15510 | 732 | |
| 34145 | 733 | lemma swap_triple: | 
| 734 | assumes "a \<noteq> c" and "b \<noteq> c" | |
| 735 | shows "swap a b (swap b c (swap a b f)) = swap a c f" | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39213diff
changeset | 736 | using assms by (simp add: fun_eq_iff swap_def) | 
| 34145 | 737 | |
| 34101 | 738 | lemma comp_swap: "f \<circ> swap a b g = swap a b (f \<circ> g)" | 
| 63322 | 739 | by (rule ext) (simp add: fun_upd_def swap_def) | 
| 34101 | 740 | |
| 39076 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 741 | lemma swap_image_eq [simp]: | 
| 63322 | 742 | assumes "a \<in> A" "b \<in> A" | 
| 743 | shows "swap a b f ` A = f ` A" | |
| 39076 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 744 | proof - | 
| 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 745 | have subset: "\<And>f. swap a b f ` A \<subseteq> f ` A" | 
| 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 746 | using assms by (auto simp: image_iff swap_def) | 
| 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 747 | then have "swap a b (swap a b f) ` A \<subseteq> (swap a b f) ` A" . | 
| 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 748 | with subset[of f] show ?thesis by auto | 
| 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 749 | qed | 
| 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 750 | |
| 63322 | 751 | lemma inj_on_imp_inj_on_swap: "inj_on f A \<Longrightarrow> a \<in> A \<Longrightarrow> b \<in> A \<Longrightarrow> inj_on (swap a b f) A" | 
| 752 | by (auto simp add: inj_on_def swap_def) | |
| 15510 | 753 | |
| 754 | lemma inj_on_swap_iff [simp]: | |
| 63322 | 755 | assumes A: "a \<in> A" "b \<in> A" | 
| 756 | shows "inj_on (swap a b f) A \<longleftrightarrow> inj_on f A" | |
| 39075 | 757 | proof | 
| 15510 | 758 | assume "inj_on (swap a b f) A" | 
| 39075 | 759 | with A have "inj_on (swap a b (swap a b f)) A" | 
| 760 | by (iprover intro: inj_on_imp_inj_on_swap) | |
| 63322 | 761 | then show "inj_on f A" by simp | 
| 15510 | 762 | next | 
| 763 | assume "inj_on f A" | |
| 63322 | 764 | with A show "inj_on (swap a b f) A" | 
| 765 | by (iprover intro: inj_on_imp_inj_on_swap) | |
| 15510 | 766 | qed | 
| 767 | ||
| 39076 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 768 | lemma surj_imp_surj_swap: "surj f \<Longrightarrow> surj (swap a b f)" | 
| 40702 | 769 | by simp | 
| 15510 | 770 | |
| 39076 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 771 | lemma surj_swap_iff [simp]: "surj (swap a b f) \<longleftrightarrow> surj f" | 
| 40702 | 772 | by simp | 
| 21547 
9c9fdf4c2949
moved order arities for fun and bool to Fun/Orderings
 haftmann parents: 
21327diff
changeset | 773 | |
| 63322 | 774 | lemma bij_betw_swap_iff [simp]: "x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> bij_betw (swap x y f) A B \<longleftrightarrow> bij_betw f A B" | 
| 39076 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 775 | by (auto simp: bij_betw_def) | 
| 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 776 | |
| 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 777 | lemma bij_swap_iff [simp]: "bij (swap a b f) \<longleftrightarrow> bij f" | 
| 
b3a9b6734663
Introduce surj_on and replace surj and bij by abbreviations.
 hoelzl parents: 
39075diff
changeset | 778 | by simp | 
| 39075 | 779 | |
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
35584diff
changeset | 780 | hide_const (open) swap | 
| 21547 
9c9fdf4c2949
moved order arities for fun and bool to Fun/Orderings
 haftmann parents: 
21327diff
changeset | 781 | |
| 56608 | 782 | |
| 60758 | 783 | subsection \<open>Inversion of injective functions\<close> | 
| 31949 | 784 | |
| 63322 | 785 | definition the_inv_into :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'a)"
 | 
| 63324 | 786 | where "the_inv_into A f = (\<lambda>x. THE y. y \<in> A \<and> f y = x)" | 
| 63322 | 787 | |
| 788 | lemma the_inv_into_f_f: "inj_on f A \<Longrightarrow> x \<in> A \<Longrightarrow> the_inv_into A f (f x) = x" | |
| 789 | unfolding the_inv_into_def inj_on_def by blast | |
| 32961 | 790 | |
| 63322 | 791 | lemma f_the_inv_into_f: "inj_on f A \<Longrightarrow> y \<in> f ` A \<Longrightarrow> f (the_inv_into A f y) = y" | 
| 792 | apply (simp add: the_inv_into_def) | |
| 793 | apply (rule the1I2) | |
| 63575 | 794 | apply (blast dest: inj_onD) | 
| 63322 | 795 | apply blast | 
| 796 | done | |
| 32961 | 797 | |
| 63322 | 798 | lemma the_inv_into_into: "inj_on f A \<Longrightarrow> x \<in> f ` A \<Longrightarrow> A \<subseteq> B \<Longrightarrow> the_inv_into A f x \<in> B" | 
| 799 | apply (simp add: the_inv_into_def) | |
| 800 | apply (rule the1I2) | |
| 63575 | 801 | apply (blast dest: inj_onD) | 
| 63322 | 802 | apply blast | 
| 803 | done | |
| 32961 | 804 | |
| 63322 | 805 | lemma the_inv_into_onto [simp]: "inj_on f A \<Longrightarrow> the_inv_into A f ` (f ` A) = A" | 
| 806 | by (fast intro: the_inv_into_into the_inv_into_f_f [symmetric]) | |
| 32961 | 807 | |
| 63322 | 808 | lemma the_inv_into_f_eq: "inj_on f A \<Longrightarrow> f x = y \<Longrightarrow> x \<in> A \<Longrightarrow> the_inv_into A f y = x" | 
| 32961 | 809 | apply (erule subst) | 
| 63322 | 810 | apply (erule the_inv_into_f_f) | 
| 811 | apply assumption | |
| 32961 | 812 | done | 
| 813 | ||
| 33057 | 814 | lemma the_inv_into_comp: | 
| 63322 | 815 | "inj_on f (g ` A) \<Longrightarrow> inj_on g A \<Longrightarrow> x \<in> f ` g ` A \<Longrightarrow> | 
| 816 | the_inv_into A (f \<circ> g) x = (the_inv_into A g \<circ> the_inv_into (g ` A) f) x" | |
| 817 | apply (rule the_inv_into_f_eq) | |
| 818 | apply (fast intro: comp_inj_on) | |
| 819 | apply (simp add: f_the_inv_into_f the_inv_into_into) | |
| 820 | apply (simp add: the_inv_into_into) | |
| 821 | done | |
| 32961 | 822 | |
| 63322 | 823 | lemma inj_on_the_inv_into: "inj_on f A \<Longrightarrow> inj_on (the_inv_into A f) (f ` A)" | 
| 824 | by (auto intro: inj_onI simp: the_inv_into_f_f) | |
| 32961 | 825 | |
| 63322 | 826 | lemma bij_betw_the_inv_into: "bij_betw f A B \<Longrightarrow> bij_betw (the_inv_into A f) B A" | 
| 827 | by (auto simp add: bij_betw_def inj_on_the_inv_into the_inv_into_into) | |
| 32961 | 828 | |
| 63322 | 829 | abbreviation the_inv :: "('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'a)"
 | 
| 830 | where "the_inv f \<equiv> the_inv_into UNIV f" | |
| 32998 
31b19fa0de0b
Renamed inv to the_inv and turned it into an abbreviation (based on the_inv_onto).
 berghofe parents: 
32988diff
changeset | 831 | |
| 64965 | 832 | lemma the_inv_f_f: "the_inv f (f x) = x" if "inj f" | 
| 833 | using that UNIV_I by (rule the_inv_into_f_f) | |
| 32998 
31b19fa0de0b
Renamed inv to the_inv and turned it into an abbreviation (based on the_inv_onto).
 berghofe parents: 
32988diff
changeset | 834 | |
| 44277 
bcb696533579
moved fundamental lemma fun_eq_iff to theory HOL; tuned whitespace
 haftmann parents: 
43991diff
changeset | 835 | |
| 60758 | 836 | subsection \<open>Cantor's Paradox\<close> | 
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 837 | |
| 63323 | 838 | theorem Cantors_paradox: "\<nexists>f. f ` A = Pow A" | 
| 839 | proof | |
| 840 | assume "\<exists>f. f ` A = Pow A" | |
| 841 | then obtain f where f: "f ` A = Pow A" .. | |
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 842 |   let ?X = "{a \<in> A. a \<notin> f a}"
 | 
| 63323 | 843 | have "?X \<in> Pow A" by blast | 
| 844 | then have "?X \<in> f ` A" by (simp only: f) | |
| 845 | then obtain x where "x \<in> A" and "f x = ?X" by blast | |
| 846 | then show False by blast | |
| 40703 
d1fc454d6735
Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
 hoelzl parents: 
40702diff
changeset | 847 | qed | 
| 31949 | 848 | |
| 63322 | 849 | |
| 61204 | 850 | subsection \<open>Setup\<close> | 
| 40969 | 851 | |
| 60758 | 852 | subsubsection \<open>Proof tools\<close> | 
| 22845 | 853 | |
| 63400 | 854 | text \<open>Simplify terms of the form \<open>f(\<dots>,x:=y,\<dots>,x:=z,\<dots>)\<close> to \<open>f(\<dots>,x:=z,\<dots>)\<close>\<close> | 
| 22845 | 855 | |
| 60758 | 856 | simproc_setup fun_upd2 ("f(v := w, x := y)") = \<open>fn _ =>
 | 
| 63322 | 857 | let | 
| 858 | fun gen_fun_upd NONE T _ _ = NONE | |
| 859 |       | gen_fun_upd (SOME f) T x y = SOME (Const (@{const_name fun_upd}, T) $ f $ x $ y)
 | |
| 860 | fun dest_fun_T1 (Type (_, T :: Ts)) = T | |
| 861 |     fun find_double (t as Const (@{const_name fun_upd},T) $ f $ x $ y) =
 | |
| 862 | let | |
| 863 |         fun find (Const (@{const_name fun_upd},T) $ g $ v $ w) =
 | |
| 864 | if v aconv x then SOME g else gen_fun_upd (find g) T v w | |
| 865 | | find t = NONE | |
| 866 | in (dest_fun_T1 T, gen_fun_upd (find f) T x y) end | |
| 24017 | 867 | |
| 63322 | 868 |     val ss = simpset_of @{context}
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51598diff
changeset | 869 | |
| 63322 | 870 | fun proc ctxt ct = | 
| 871 | let | |
| 872 | val t = Thm.term_of ct | |
| 873 | in | |
| 63400 | 874 | (case find_double t of | 
| 63322 | 875 | (T, NONE) => NONE | 
| 876 | | (T, SOME rhs) => | |
| 877 | SOME (Goal.prove ctxt [] [] (Logic.mk_equals (t, rhs)) | |
| 878 | (fn _ => | |
| 879 | resolve_tac ctxt [eq_reflection] 1 THEN | |
| 880 |                 resolve_tac ctxt @{thms ext} 1 THEN
 | |
| 63400 | 881 | simp_tac (put_simpset ss ctxt) 1))) | 
| 63322 | 882 | end | 
| 883 | in proc end | |
| 60758 | 884 | \<close> | 
| 22845 | 885 | |
| 886 | ||
| 60758 | 887 | subsubsection \<open>Functorial structure of types\<close> | 
| 40969 | 888 | |
| 55467 
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
 blanchet parents: 
55414diff
changeset | 889 | ML_file "Tools/functor.ML" | 
| 40969 | 890 | |
| 55467 
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
 blanchet parents: 
55414diff
changeset | 891 | functor map_fun: map_fun | 
| 47488 
be6dd389639d
centralized enriched_type declaration, thanks to in-situ available Isar commands
 haftmann parents: 
46950diff
changeset | 892 | by (simp_all add: fun_eq_iff) | 
| 
be6dd389639d
centralized enriched_type declaration, thanks to in-situ available Isar commands
 haftmann parents: 
46950diff
changeset | 893 | |
| 55467 
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
 blanchet parents: 
55414diff
changeset | 894 | functor vimage | 
| 49739 | 895 | by (simp_all add: fun_eq_iff vimage_comp) | 
| 896 | ||
| 63322 | 897 | |
| 60758 | 898 | text \<open>Legacy theorem names\<close> | 
| 49739 | 899 | |
| 900 | lemmas o_def = comp_def | |
| 901 | lemmas o_apply = comp_apply | |
| 902 | lemmas o_assoc = comp_assoc [symmetric] | |
| 903 | lemmas id_o = id_comp | |
| 904 | lemmas o_id = comp_id | |
| 905 | lemmas o_eq_dest = comp_eq_dest | |
| 906 | lemmas o_eq_elim = comp_eq_elim | |
| 55066 | 907 | lemmas o_eq_dest_lhs = comp_eq_dest_lhs | 
| 908 | lemmas o_eq_id_dest = comp_eq_id_dest | |
| 47488 
be6dd389639d
centralized enriched_type declaration, thanks to in-situ available Isar commands
 haftmann parents: 
46950diff
changeset | 909 | |
| 2912 | 910 | end |