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