# HG changeset patch # User hoelzl # Date 1290762244 -3600 # Node ID 407c6122956fe4f9cfbf3b055a2c94831c0b9841 # Parent e7aa34600c361ff4fcc6c9ef4b66b5992d56eab8# Parent d1fc454d67357820e94f1ec58f25f9d89df3d40f merged diff -r e7aa34600c36 -r 407c6122956f NEWS --- a/NEWS Fri Nov 26 09:15:49 2010 +0100 +++ b/NEWS Fri Nov 26 10:04:04 2010 +0100 @@ -291,9 +291,10 @@ derive instantiated and simplified equations for inductive predicates, similar to inductive_cases. -* "bij f" is now an abbreviation of "bij_betw f UNIV UNIV". surj_on is a -generalization of surj, and "surj f" is now a abbreviation of "surj_on f UNIV". -The theorems bij_def and surj_def are unchanged. +* "bij f" is now an abbreviation of "bij_betw f UNIV UNIV". "surj f" is now an +abbreviation of "range f = UIV". The theorems bij_def and surj_def are +unchanged. +INCOMPATIBILITY. * Function package: .psimps rules are no longer implicitly declared [simp]. INCOMPATIBILITY. diff -r e7aa34600c36 -r 407c6122956f src/HOL/Finite_Set.thy --- a/src/HOL/Finite_Set.thy Fri Nov 26 09:15:49 2010 +0100 +++ b/src/HOL/Finite_Set.thy Fri Nov 26 10:04:04 2010 +0100 @@ -2179,6 +2179,11 @@ finite A; finite B |] ==> card A = card B" by (auto intro: le_antisym card_inj_on_le) +lemma bij_betw_finite: + assumes "bij_betw f A B" + shows "finite A \ finite B" +using assms unfolding bij_betw_def +using finite_imageD[of f A] by auto subsubsection {* Pigeonhole Principles *} @@ -2286,7 +2291,7 @@ lemma finite_UNIV_surj_inj: fixes f :: "'a \ 'a" shows "finite(UNIV:: 'a set) \ surj f \ inj f" -by (blast intro: finite_surj_inj subset_UNIV dest:surj_range) +by (blast intro: finite_surj_inj subset_UNIV) lemma finite_UNIV_inj_surj: fixes f :: "'a \ 'a" shows "finite(UNIV:: 'a set) \ inj f \ surj f" diff -r e7aa34600c36 -r 407c6122956f src/HOL/Fun.thy --- a/src/HOL/Fun.thy Fri Nov 26 09:15:49 2010 +0100 +++ b/src/HOL/Fun.thy Fri Nov 26 10:04:04 2010 +0100 @@ -130,24 +130,22 @@ by (simp_all add: fun_eq_iff) -subsection {* Injectivity, Surjectivity and Bijectivity *} +subsection {* Injectivity and Bijectivity *} definition inj_on :: "('a \ 'b) \ 'a set \ bool" where -- "injective" "inj_on f A \ (\x\A. \y\A. f x = f y \ x = y)" -definition surj_on :: "('a \ 'b) \ 'b set \ bool" where -- "surjective" - "surj_on f B \ B \ range f" - definition bij_betw :: "('a \ 'b) \ 'a set \ 'b set \ bool" where -- "bijective" "bij_betw f A B \ inj_on f A \ f ` A = B" -text{*A common special case: functions injective over the entire domain type.*} +text{*A common special case: functions injective, surjective or bijective over +the entire domain type.*} abbreviation "inj f \ inj_on f UNIV" -abbreviation - "surj f \ surj_on f UNIV" +abbreviation surj :: "('a \ 'b) \ bool" where -- "surjective" + "surj f \ (range f = UNIV)" abbreviation "bij f \ bij_betw f UNIV UNIV" @@ -171,6 +169,14 @@ lemma inj_on_eq_iff: "inj_on f A ==> x:A ==> y:A ==> (f(x) = f(y)) = (x=y)" by (force simp add: inj_on_def) +lemma inj_on_cong: + "(\ a. a : A \ f a = g a) \ inj_on f A = inj_on g A" +unfolding inj_on_def by auto + +lemma inj_on_strict_subset: + "\ inj_on f B; A < B \ \ f`A < f`B" +unfolding inj_on_def unfolding image_def by blast + lemma inj_comp: "inj f \ inj g \ inj (f \ g)" by (simp add: inj_on_def) @@ -187,8 +193,44 @@ lemma inj_on_id2[simp]: "inj_on (%x. x) A" by (simp add: inj_on_def) -lemma surj_id[simp]: "surj_on id A" -by (simp add: surj_on_def) +lemma inj_on_Int: "\inj_on f A; inj_on f B\ \ inj_on f (A \ B)" +unfolding inj_on_def by blast + +lemma inj_on_INTER: + "\I \ {}; \ i. i \ I \ inj_on f (A i)\ \ inj_on f (\ i \ I. A i)" +unfolding inj_on_def by blast + +lemma inj_on_Inter: + "\S \ {}; \ A. A \ S \ inj_on f A\ \ inj_on f (Inter S)" +unfolding inj_on_def by blast + +lemma inj_on_UNION_chain: + assumes CH: "\ i j. \i \ I; j \ I\ \ A i \ A j \ A j \ A i" and + INJ: "\ i. i \ I \ inj_on f (A i)" + shows "inj_on f (\ i \ I. A i)" +proof(unfold inj_on_def UNION_def, auto) + fix i j x y + assume *: "i \ I" "j \ I" and **: "x \ A i" "y \ A j" + and ***: "f x = f y" + show "x = y" + proof- + {assume "A i \ A j" + with ** have "x \ A j" by auto + with INJ * ** *** have ?thesis + by(auto simp add: inj_on_def) + } + moreover + {assume "A j \ A i" + with ** have "y \ A i" by auto + with INJ * ** *** have ?thesis + by(auto simp add: inj_on_def) + } + ultimately show ?thesis using CH * by blast + qed +qed + +lemma surj_id: "surj id" +by simp lemma bij_id[simp]: "bij id" by (simp add: bij_betw_def) @@ -251,20 +293,19 @@ apply (blast) done -lemma surj_onI: "(\x. x \ B \ g (f x) = x) \ surj_on g B" - by (simp add: surj_on_def) (blast intro: sym) +lemma comp_inj_on_iff: + "inj_on f A \ inj_on f' (f ` A) \ inj_on (f' o f) A" +by(auto simp add: comp_inj_on inj_on_def) -lemma surj_onD: "surj_on f B \ y \ B \ \x. y = f x" - by (auto simp: surj_on_def) - -lemma surj_on_range_iff: "surj_on f B \ (\A. f ` A = B)" - unfolding surj_on_def by (auto intro!: exI[of _ "f -` B"]) +lemma inj_on_imageI2: + "inj_on (f' o f) A \ inj_on f A" +by(auto simp add: comp_inj_on inj_on_def) lemma surj_def: "surj f \ (\y. \x. y = f x)" - by (simp add: surj_on_def subset_eq image_iff) + by auto -lemma surjI: "(\ x. g (f x) = x) \ surj g" - by (blast intro: surj_onI) +lemma surjI: assumes *: "\ x. g (f x) = x" shows "surj g" + using *[symmetric] by auto lemma surjD: "surj f \ \x. y = f x" by (simp add: surj_def) @@ -278,17 +319,25 @@ apply (drule_tac x = x in spec, blast) done -lemma surj_range: "surj f \ range f = UNIV" - by (auto simp add: surj_on_def) +lemma bij_betw_imp_surj: "bij_betw f A UNIV \ surj f" + unfolding bij_betw_def by auto + +lemma bij_betw_empty1: + assumes "bij_betw f {} A" + shows "A = {}" +using assms unfolding bij_betw_def by blast -lemma surj_range_iff: "surj f \ range f = UNIV" - unfolding surj_on_def by auto +lemma bij_betw_empty2: + assumes "bij_betw f A {}" + shows "A = {}" +using assms unfolding bij_betw_def by blast -lemma bij_betw_imp_surj: "bij_betw f A UNIV \ surj f" - unfolding bij_betw_def surj_range_iff by auto +lemma inj_on_imp_bij_betw: + "inj_on f A \ bij_betw f A (f ` A)" +unfolding bij_betw_def by simp lemma bij_def: "bij f \ inj f \ surj f" - unfolding surj_range_iff bij_betw_def .. + unfolding bij_betw_def .. lemma bijI: "[| inj f; surj f |] ==> bij f" by (simp add: bij_def) @@ -302,16 +351,39 @@ lemma bij_betw_imp_inj_on: "bij_betw f A B \ inj_on f A" by (simp add: bij_betw_def) -lemma bij_betw_imp_surj_on: "bij_betw f A B \ surj_on f B" -by (auto simp: bij_betw_def surj_on_range_iff) - -lemma bij_comp: "bij f \ bij g \ bij (g o f)" -by(fastsimp intro: comp_inj_on comp_surj simp: bij_def surj_range) - lemma bij_betw_trans: "bij_betw f A B \ bij_betw g B C \ bij_betw (g o f) A C" by(auto simp add:bij_betw_def comp_inj_on) +lemma bij_comp: "bij f \ bij g \ bij (g o f)" + by (rule bij_betw_trans) + +lemma bij_betw_comp_iff: + "bij_betw f A A' \ bij_betw f' A' A'' \ bij_betw (f' o f) A A''" +by(auto simp add: bij_betw_def inj_on_def) + +lemma bij_betw_comp_iff2: + assumes BIJ: "bij_betw f' A' A''" and IM: "f ` A \ A'" + shows "bij_betw f A A' \ bij_betw (f' o f) A A''" +using assms +proof(auto simp add: bij_betw_comp_iff) + assume *: "bij_betw (f' \ f) A A''" + thus "bij_betw f A A'" + using IM + proof(auto simp add: bij_betw_def) + assume "inj_on (f' \ f) A" + thus "inj_on f A" using inj_on_imageI2 by blast + next + fix a' assume **: "a' \ A'" + hence "f' a' \ A''" using BIJ unfolding bij_betw_def by auto + then obtain a where 1: "a \ A \ f'(f a) = f' a'" using * + unfolding bij_betw_def by force + hence "f a \ A'" using IM by auto + hence "f a = a'" using BIJ ** 1 unfolding bij_betw_def inj_on_def by auto + thus "a' \ f ` A" using 1 by auto + qed +qed + lemma bij_betw_inv: assumes "bij_betw f A B" shows "EX g. bij_betw g B A" proof - have i: "inj_on f A" and s: "f ` A = B" @@ -343,21 +415,81 @@ ultimately show ?thesis by(auto simp:bij_betw_def) qed +lemma bij_betw_cong: + "(\ a. a \ A \ f a = g a) \ bij_betw f A A' = bij_betw g A A'" +unfolding bij_betw_def inj_on_def by force + +lemma bij_betw_id[intro, simp]: + "bij_betw id A A" +unfolding bij_betw_def id_def by auto + +lemma bij_betw_id_iff: + "bij_betw id A B \ A = B" +by(auto simp add: bij_betw_def) + lemma bij_betw_combine: assumes "bij_betw f A B" "bij_betw f C D" "B \ D = {}" shows "bij_betw f (A \ C) (B \ D)" using assms unfolding bij_betw_def inj_on_Un image_Un by auto +lemma bij_betw_UNION_chain: + assumes CH: "\ i j. \i \ I; j \ I\ \ A i \ A j \ A j \ A i" and + BIJ: "\ i. i \ I \ bij_betw f (A i) (A' i)" + shows "bij_betw f (\ i \ I. A i) (\ i \ I. A' i)" +proof(unfold bij_betw_def, auto simp add: image_def) + have "\ i. i \ I \ inj_on f (A i)" + using BIJ bij_betw_def[of f] by auto + thus "inj_on f (\ i \ I. A i)" + using CH inj_on_UNION_chain[of I A f] by auto +next + fix i x + assume *: "i \ I" "x \ A i" + hence "f x \ A' i" using BIJ bij_betw_def[of f] by auto + thus "\j \ I. f x \ A' j" using * by blast +next + fix i x' + assume *: "i \ I" "x' \ A' i" + hence "\x \ A i. x' = f x" using BIJ bij_betw_def[of f] by blast + thus "\j \ I. \x \ A j. x' = f x" + using * by blast +qed + +lemma bij_betw_Disj_Un: + assumes DISJ: "A \ B = {}" and DISJ': "A' \ B' = {}" and + B1: "bij_betw f A A'" and B2: "bij_betw f B B'" + shows "bij_betw f (A \ B) (A' \ B')" +proof- + have 1: "inj_on f A \ inj_on f B" + using B1 B2 by (auto simp add: bij_betw_def) + have 2: "f`A = A' \ f`B = B'" + using B1 B2 by (auto simp add: bij_betw_def) + hence "f`(A - B) \ f`(B - A) = {}" + using DISJ DISJ' by blast + hence "inj_on f (A \ B)" + using 1 by (auto simp add: inj_on_Un) + (* *) + moreover + have "f`(A \ B) = A' \ B'" + using 2 by auto + ultimately show ?thesis + unfolding bij_betw_def by auto +qed + +lemma bij_betw_subset: + assumes BIJ: "bij_betw f A A'" and + SUB: "B \ A" and IM: "f ` B = B'" + shows "bij_betw f B B'" +using assms +by(unfold bij_betw_def inj_on_def, auto simp add: inj_on_def) + lemma surj_image_vimage_eq: "surj f ==> f ` (f -` A) = A" -by (simp add: surj_range) +by simp lemma inj_vimage_image_eq: "inj f ==> f -` (f ` A) = A" by (simp add: inj_on_def, blast) lemma vimage_subsetD: "surj f ==> f -` B <= A ==> B <= f ` A" -apply (unfold surj_def) -apply (blast intro: sym) -done +by (blast intro: sym) lemma vimage_subsetI: "inj f ==> B <= f ` A ==> f -` B <= A" by (unfold inj_on_def, blast) @@ -410,7 +542,7 @@ done lemma surj_Compl_image_subset: "surj f ==> -(f`A) <= f`(-A)" -by (auto simp add: surj_def) +by auto lemma inj_image_Compl_subset: "inj f ==> f`(-A) <= -(f`A)" by (auto simp add: inj_on_def) @@ -559,10 +691,10 @@ qed lemma surj_imp_surj_swap: "surj f \ surj (swap a b f)" - unfolding surj_range_iff by simp + by simp lemma surj_swap_iff [simp]: "surj (swap a b f) \ surj f" - unfolding surj_range_iff by simp + by simp lemma bij_betw_swap_iff [simp]: "\ x \ A; y \ A \ \ bij_betw (swap x y f) A B \ bij_betw f A B" @@ -635,6 +767,17 @@ shows "the_inv f (f x) = x" using assms UNIV_I by (rule the_inv_into_f_f) +subsection {* Cantor's Paradox *} + +lemma Cantors_paradox: + "\(\f. f ` A = Pow A)" +proof clarify + fix f assume "f ` A = Pow A" hence *: "Pow A \ f ` A" by blast + let ?X = "{a \ A. a \ f a}" + have "?X \ Pow A" unfolding Pow_def by auto + with * obtain x where "x \ A \ f x = ?X" by blast + thus False by best +qed subsection {* Proof tool setup *} diff -r e7aa34600c36 -r 407c6122956f src/HOL/Hilbert_Choice.thy --- a/src/HOL/Hilbert_Choice.thy Fri Nov 26 09:15:49 2010 +0100 +++ b/src/HOL/Hilbert_Choice.thy Fri Nov 26 10:04:04 2010 +0100 @@ -151,10 +151,10 @@ by(fastsimp simp: image_def) lemma inj_imp_surj_inv: "inj f ==> surj (inv f)" -by (blast intro: surjI inv_into_f_f) +by (blast intro!: surjI inv_into_f_f) lemma surj_f_inv_f: "surj f ==> f(inv f y) = y" -by (simp add: f_inv_into_f surj_range) +by (simp add: f_inv_into_f) lemma inv_into_injective: assumes eq: "inv_into A f x = inv_into A f y" @@ -173,12 +173,13 @@ by (auto simp add: bij_betw_def inj_on_inv_into) lemma surj_imp_inj_inv: "surj f ==> inj (inv f)" -by (simp add: inj_on_inv_into surj_range) +by (simp add: inj_on_inv_into) lemma surj_iff: "(surj f) = (f o inv f = id)" -apply (simp add: o_def fun_eq_iff) -apply (blast intro: surjI surj_f_inv_f) -done +by (auto intro!: surjI simp: surj_f_inv_f fun_eq_iff[where 'b='a]) + +lemma surj_iff_all: "surj f \ (\x. f (inv f x) = x)" + unfolding surj_iff by (simp add: o_def fun_eq_iff) lemma surj_imp_inv_eq: "[| surj f; \x. g(f x) = x |] ==> inv f = g" apply (rule ext) @@ -260,6 +261,208 @@ ultimately show "finite (UNIV :: 'a set)" by simp qed +lemma image_inv_into_cancel: + assumes SURJ: "f`A=A'" and SUB: "B' \ A'" + shows "f `((inv_into A f)`B') = B'" + using assms +proof (auto simp add: f_inv_into_f) + let ?f' = "(inv_into A f)" + fix a' assume *: "a' \ B'" + then have "a' \ A'" using SUB by auto + then have "a' = f (?f' a')" + using SURJ by (auto simp add: f_inv_into_f) + then show "a' \ f ` (?f' ` B')" using * by blast +qed + +lemma inv_into_inv_into_eq: + assumes "bij_betw f A A'" "a \ A" + shows "inv_into A' (inv_into A f) a = f a" +proof - + let ?f' = "inv_into A f" let ?f'' = "inv_into A' ?f'" + have 1: "bij_betw ?f' A' A" using assms + by (auto simp add: bij_betw_inv_into) + obtain a' where 2: "a' \ A'" and 3: "?f' a' = a" + using 1 `a \ A` unfolding bij_betw_def by force + hence "?f'' a = a'" + using `a \ A` 1 3 by (auto simp add: f_inv_into_f bij_betw_def) + moreover have "f a = a'" using assms 2 3 + by (auto simp add: inv_into_f_f bij_betw_def) + ultimately show "?f'' a = f a" by simp +qed + +lemma inj_on_iff_surj: + assumes "A \ {}" + shows "(\f. inj_on f A \ f ` A \ A') \ (\g. g ` A' = A)" +proof safe + fix f assume INJ: "inj_on f A" and INCL: "f ` A \ A'" + let ?phi = "\a' a. a \ A \ f a = a'" let ?csi = "\a. a \ A" + let ?g = "\a'. if a' \ f ` A then (SOME a. ?phi a' a) else (SOME a. ?csi a)" + have "?g ` A' = A" + proof + show "?g ` A' \ A" + proof clarify + fix a' assume *: "a' \ A'" + show "?g a' \ A" + proof cases + assume Case1: "a' \ f ` A" + then obtain a where "?phi a' a" by blast + hence "?phi a' (SOME a. ?phi a' a)" using someI[of "?phi a'" a] by blast + with Case1 show ?thesis by auto + next + assume Case2: "a' \ f ` A" + hence "?csi (SOME a. ?csi a)" using assms someI_ex[of ?csi] by blast + with Case2 show ?thesis by auto + qed + qed + next + show "A \ ?g ` A'" + proof- + {fix a assume *: "a \ A" + let ?b = "SOME aa. ?phi (f a) aa" + have "?phi (f a) a" using * by auto + hence 1: "?phi (f a) ?b" using someI[of "?phi(f a)" a] by blast + hence "?g(f a) = ?b" using * by auto + moreover have "a = ?b" using 1 INJ * by (auto simp add: inj_on_def) + ultimately have "?g(f a) = a" by simp + with INCL * have "?g(f a) = a \ f a \ A'" by auto + } + thus ?thesis by force + qed + qed + thus "\g. g ` A' = A" by blast +next + fix g let ?f = "inv_into A' g" + have "inj_on ?f (g ` A')" + by (auto simp add: inj_on_inv_into) + moreover + {fix a' assume *: "a' \ A'" + let ?phi = "\ b'. b' \ A' \ g b' = g a'" + have "?phi a'" using * by auto + hence "?phi(SOME b'. ?phi b')" using someI[of ?phi] by blast + hence "?f(g a') \ A'" unfolding inv_into_def by auto + } + ultimately show "\f. inj_on f (g ` A') \ f ` g ` A' \ A'" by auto +qed + +lemma Ex_inj_on_UNION_Sigma: + "\f. (inj_on f (\ i \ I. A i) \ f ` (\ i \ I. A i) \ (SIGMA i : I. A i))" +proof + let ?phi = "\ a i. i \ I \ a \ A i" + let ?sm = "\ a. SOME i. ?phi a i" + let ?f = "\a. (?sm a, a)" + have "inj_on ?f (\ i \ I. A i)" unfolding inj_on_def by auto + moreover + { { fix i a assume "i \ I" and "a \ A i" + hence "?sm a \ I \ a \ A(?sm a)" using someI[of "?phi a" i] by auto + } + hence "?f ` (\ i \ I. A i) \ (SIGMA i : I. A i)" by auto + } + ultimately + show "inj_on ?f (\ i \ I. A i) \ ?f ` (\ i \ I. A i) \ (SIGMA i : I. A i)" + by auto +qed + +subsection {* The Cantor-Bernstein Theorem *} + +lemma Cantor_Bernstein_aux: + shows "\A' h. A' \ A \ + (\a \ A'. a \ g`(B - f ` A')) \ + (\a \ A'. h a = f a) \ + (\a \ A - A'. h a \ B - (f ` A') \ a = g(h a))" +proof- + obtain H where H_def: "H = (\ A'. A - (g`(B - (f ` A'))))" by blast + have 0: "mono H" unfolding mono_def H_def by blast + then obtain A' where 1: "H A' = A'" using lfp_unfold by blast + hence 2: "A' = A - (g`(B - (f ` A')))" unfolding H_def by simp + hence 3: "A' \ A" by blast + have 4: "\a \ A'. a \ g`(B - f ` A')" + using 2 by blast + have 5: "\a \ A - A'. \b \ B - (f ` A'). a = g b" + using 2 by blast + (* *) + obtain h where h_def: + "h = (\ a. if a \ A' then f a else (SOME b. b \ B - (f ` A') \ a = g b))" by blast + hence "\a \ A'. h a = f a" by auto + moreover + have "\a \ A - A'. h a \ B - (f ` A') \ a = g(h a)" + proof + fix a assume *: "a \ A - A'" + let ?phi = "\ b. b \ B - (f ` A') \ a = g b" + have "h a = (SOME b. ?phi b)" using h_def * by auto + moreover have "\b. ?phi b" using 5 * by auto + ultimately show "?phi (h a)" using someI_ex[of ?phi] by auto + qed + ultimately show ?thesis using 3 4 by blast +qed + +theorem Cantor_Bernstein: + assumes INJ1: "inj_on f A" and SUB1: "f ` A \ B" and + INJ2: "inj_on g B" and SUB2: "g ` B \ A" + shows "\h. bij_betw h A B" +proof- + obtain A' and h where 0: "A' \ A" and + 1: "\a \ A'. a \ g`(B - f ` A')" and + 2: "\a \ A'. h a = f a" and + 3: "\a \ A - A'. h a \ B - (f ` A') \ a = g(h a)" + using Cantor_Bernstein_aux[of A g B f] by blast + have "inj_on h A" + proof (intro inj_onI) + fix a1 a2 + assume 4: "a1 \ A" and 5: "a2 \ A" and 6: "h a1 = h a2" + show "a1 = a2" + proof(cases "a1 \ A'") + assume Case1: "a1 \ A'" + show ?thesis + proof(cases "a2 \ A'") + assume Case11: "a2 \ A'" + hence "f a1 = f a2" using Case1 2 6 by auto + thus ?thesis using INJ1 Case1 Case11 0 + unfolding inj_on_def by blast + next + assume Case12: "a2 \ A'" + hence False using 3 5 2 6 Case1 by force + thus ?thesis by simp + qed + next + assume Case2: "a1 \ A'" + show ?thesis + proof(cases "a2 \ A'") + assume Case21: "a2 \ A'" + hence False using 3 4 2 6 Case2 by auto + thus ?thesis by simp + next + assume Case22: "a2 \ A'" + hence "a1 = g(h a1) \ a2 = g(h a2)" using Case2 4 5 3 by auto + thus ?thesis using 6 by simp + qed + qed + qed + (* *) + moreover + have "h ` A = B" + proof safe + fix a assume "a \ A" + thus "h a \ B" using SUB1 2 3 by (case_tac "a \ A'", auto) + next + fix b assume *: "b \ B" + show "b \ h ` A" + proof(cases "b \ f ` A'") + assume Case1: "b \ f ` A'" + then obtain a where "a \ A' \ b = f a" by blast + thus ?thesis using 2 0 by force + next + assume Case2: "b \ f ` A'" + hence "g b \ A'" using 1 * by auto + hence 4: "g b \ A - A'" using * SUB2 by auto + hence "h(g b) \ B \ g(h(g b)) = g b" + using 3 by auto + hence "h(g b) = b" using * INJ2 unfolding inj_on_def by auto + thus ?thesis using 4 by force + qed + qed + (* *) + ultimately show ?thesis unfolding bij_betw_def by auto +qed subsection {*Other Consequences of Hilbert's Epsilon*} diff -r e7aa34600c36 -r 407c6122956f src/HOL/Lattice/Orders.thy --- a/src/HOL/Lattice/Orders.thy Fri Nov 26 09:15:49 2010 +0100 +++ b/src/HOL/Lattice/Orders.thy Fri Nov 26 10:04:04 2010 +0100 @@ -118,8 +118,8 @@ qed qed -lemma range_dual [simp]: "dual ` UNIV = UNIV" -proof (rule surj_range) +lemma range_dual [simp]: "surj dual" +proof - have "\x'. dual (undual x') = x'" by simp thus "surj dual" by (rule surjI) qed diff -r e7aa34600c36 -r 407c6122956f src/HOL/Library/ContNotDenum.thy --- a/src/HOL/Library/ContNotDenum.thy Fri Nov 26 09:15:49 2010 +0100 +++ b/src/HOL/Library/ContNotDenum.thy Fri Nov 26 10:04:04 2010 +0100 @@ -565,8 +565,7 @@ shows "\ (\f::nat\real. surj f)" proof -- "by contradiction" assume "\f::nat\real. surj f" - then obtain f::"nat\real" where "surj f" by auto - hence rangeF: "range f = UNIV" by (rule surj_range) + then obtain f::"nat\real" where rangeF: "surj f" by auto -- "We now produce a real number x that is not in the range of f, using the properties of newInt. " have "\x. x \ (\n. newInt n f)" using newInt_notempty by blast moreover have "\n. f n \ (\n. newInt n f)" by (rule newInt_inter) diff -r e7aa34600c36 -r 407c6122956f src/HOL/Library/Countable.thy --- a/src/HOL/Library/Countable.thy Fri Nov 26 09:15:49 2010 +0100 +++ b/src/HOL/Library/Countable.thy Fri Nov 26 10:04:04 2010 +0100 @@ -165,7 +165,7 @@ qed lemma Rats_eq_range_nat_to_rat_surj: "\ = range nat_to_rat_surj" -by (simp add: Rats_def surj_nat_to_rat_surj surj_range) +by (simp add: Rats_def surj_nat_to_rat_surj) context field_char_0 begin diff -r e7aa34600c36 -r 407c6122956f src/HOL/Multivariate_Analysis/Derivative.thy --- a/src/HOL/Multivariate_Analysis/Derivative.thy Fri Nov 26 09:15:49 2010 +0100 +++ b/src/HOL/Multivariate_Analysis/Derivative.thy Fri Nov 26 10:04:04 2010 +0100 @@ -948,14 +948,12 @@ assumes lf: "linear f" and gf: "f o g = id" shows "linear g" proof- - from gf have fi: "surj f" apply (auto simp add: surj_def o_def id_def fun_eq_iff) - by metis + from gf have fi: "surj f" by (auto simp add: surj_def o_def id_def) metis from linear_surjective_isomorphism[OF lf fi] obtain h:: "'a => 'a" where h: "linear h" "\x. h (f x) = x" "\x. f (h x) = x" by blast have "h = g" apply (rule ext) using gf h(2,3) - apply (simp add: o_def id_def fun_eq_iff) - by metis + by (simp add: o_def id_def fun_eq_iff) metis with h(1) show ?thesis by blast qed diff -r e7aa34600c36 -r 407c6122956f src/HOL/Multivariate_Analysis/Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Fri Nov 26 09:15:49 2010 +0100 +++ b/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Fri Nov 26 10:04:04 2010 +0100 @@ -2843,7 +2843,7 @@ h: "linear h" "\ x\ basis ` {..i 'a" assume lf: "linear f" "linear f'" and f: "f o f' = id" from f have sf: "surj f" - apply (auto simp add: o_def fun_eq_iff id_def surj_def) + apply (auto simp add: o_def id_def surj_def) by metis from linear_surjective_isomorphism[OF lf(1) sf] lf f have "f' o f = id" unfolding fun_eq_iff o_def id_def diff -r e7aa34600c36 -r 407c6122956f src/HOL/Nominal/Examples/Support.thy --- a/src/HOL/Nominal/Examples/Support.thy Fri Nov 26 09:15:49 2010 +0100 +++ b/src/HOL/Nominal/Examples/Support.thy Fri Nov 26 10:04:04 2010 +0100 @@ -47,7 +47,7 @@ also have "\ = (\n. atom n) ` ({n. \i. n = 2*i \ n = 2*i+1})" by auto also have "\ = (\n. atom n) ` (UNIV::nat set)" using even_or_odd by auto also have "\ = (UNIV::atom set)" using atom.exhaust - by (rule_tac surj_range) (auto simp add: surj_def) + by (auto simp add: surj_def) finally show "EVEN \ ODD = UNIV" by simp qed diff -r e7aa34600c36 -r 407c6122956f src/HOL/Probability/Sigma_Algebra.thy --- a/src/HOL/Probability/Sigma_Algebra.thy Fri Nov 26 09:15:49 2010 +0100 +++ b/src/HOL/Probability/Sigma_Algebra.thy Fri Nov 26 10:04:04 2010 +0100 @@ -132,7 +132,7 @@ by (auto intro!: exI[of _ "from_nat i"]) qed have **: "range ?A' = range A" - using surj_range[OF surj_from_nat] + using surj_from_nat by (auto simp: image_compose intro!: imageI) show ?thesis unfolding * ** .. qed diff -r e7aa34600c36 -r 407c6122956f src/HOL/Product_Type.thy --- a/src/HOL/Product_Type.thy Fri Nov 26 09:15:49 2010 +0100 +++ b/src/HOL/Product_Type.thy Fri Nov 26 10:04:04 2010 +0100 @@ -1135,6 +1135,7 @@ qed lemma map_pair_surj: + fixes f :: "'a \ 'b" and g :: "'c \ 'd" assumes "surj f" and "surj g" shows "surj (map_pair f g)" unfolding surj_def diff -r e7aa34600c36 -r 407c6122956f src/HOL/Set.thy --- a/src/HOL/Set.thy Fri Nov 26 09:15:49 2010 +0100 +++ b/src/HOL/Set.thy Fri Nov 26 10:04:04 2010 +0100 @@ -622,6 +622,8 @@ lemma Pow_top: "A \ Pow A" by simp +lemma Pow_not_empty: "Pow A \ {}" + using Pow_top by blast subsubsection {* Set complement *} @@ -972,6 +974,21 @@ lemmas [symmetric, rulify] = atomize_ball and [symmetric, defn] = atomize_ball +lemma image_Pow_mono: + assumes "f ` A \ B" + shows "(image f) ` (Pow A) \ Pow B" +using assms by blast + +lemma image_Pow_surj: + assumes "f ` A = B" + shows "(image f) ` (Pow A) = Pow B" +using assms unfolding Pow_def proof(auto) + fix Y assume *: "Y \ f ` A" + obtain X where X_def: "X = {x \ A. f x \ Y}" by blast + have "f ` X = Y \ X \ A" unfolding X_def using * by auto + thus "Y \ (image f) ` {X. X \ A}" by blast +qed + subsubsection {* Derived rules involving subsets. *} text {* @{text insert}. *} diff -r e7aa34600c36 -r 407c6122956f src/HOL/SetInterval.thy --- a/src/HOL/SetInterval.thy Fri Nov 26 09:15:49 2010 +0100 +++ b/src/HOL/SetInterval.thy Fri Nov 26 10:04:04 2010 +0100 @@ -159,6 +159,10 @@ apply simp_all done +lemma lessThan_strict_subset_iff: + fixes m n :: "'a::linorder" + shows "{.. m < n" + by (metis leD lessThan_subset_iff linorder_linear not_less_iff_gr_or_eq psubset_eq) subsection {*Two-sided intervals*} @@ -262,6 +266,18 @@ apply (simp add: less_imp_le) done +lemma atLeastLessThan_inj: + fixes a b c d :: "'a::linorder" + assumes eq: "{a ..< b} = {c ..< d}" and "a < b" "c < d" + shows "a = c" "b = d" +using assms by (metis atLeastLessThan_subset_iff eq less_le_not_le linorder_antisym_conv2 subset_refl)+ + +lemma atLeastLessThan_eq_iff: + fixes a b c d :: "'a::linorder" + assumes "a < b" "c < d" + shows "{a ..< b} = {c ..< d} \ a = c \ b = d" + using atLeastLessThan_inj assms by auto + subsubsection {* Intersection *} context linorder @@ -705,6 +721,39 @@ "finite M \ \h. bij_betw h {1 .. card M} M" by (rule finite_same_card_bij) auto +lemma bij_betw_iff_card: + assumes FIN: "finite A" and FIN': "finite B" + shows BIJ: "(\f. bij_betw f A B) \ (card A = card B)" +using assms +proof(auto simp add: bij_betw_same_card) + assume *: "card A = card B" + obtain f where "bij_betw f A {0 ..< card A}" + using FIN ex_bij_betw_finite_nat by blast + moreover obtain g where "bij_betw g {0 ..< card B} B" + using FIN' ex_bij_betw_nat_finite by blast + ultimately have "bij_betw (g o f) A B" + using * by (auto simp add: bij_betw_trans) + thus "(\f. bij_betw f A B)" by blast +qed + +lemma inj_on_iff_card_le: + assumes FIN: "finite A" and FIN': "finite B" + shows "(\f. inj_on f A \ f ` A \ B) = (card A \ card B)" +proof (safe intro!: card_inj_on_le) + assume *: "card A \ card B" + obtain f where 1: "inj_on f A" and 2: "f ` A = {0 ..< card A}" + using FIN ex_bij_betw_finite_nat unfolding bij_betw_def by force + moreover obtain g where "inj_on g {0 ..< card B}" and 3: "g ` {0 ..< card B} = B" + using FIN' ex_bij_betw_nat_finite unfolding bij_betw_def by force + ultimately have "inj_on g (f ` A)" using subset_inj_on[of g _ "f ` A"] * by force + hence "inj_on (g o f) A" using 1 comp_inj_on by blast + moreover + {have "{0 ..< card A} \ {0 ..< card B}" using * by force + with 2 have "f ` A \ {0 ..< card B}" by blast + hence "(g o f) ` A \ B" unfolding comp_def using 3 by force + } + ultimately show "(\f. inj_on f A \ f ` A \ B)" by blast +qed (insert assms, auto) subsection {* Intervals of integers *} diff -r e7aa34600c36 -r 407c6122956f src/HOL/UNITY/Comp/Alloc.thy --- a/src/HOL/UNITY/Comp/Alloc.thy Fri Nov 26 09:15:49 2010 +0100 +++ b/src/HOL/UNITY/Comp/Alloc.thy Fri Nov 26 10:04:04 2010 +0100 @@ -358,7 +358,7 @@ done lemma surj_sysOfAlloc [iff]: "surj sysOfAlloc" - apply (simp add: surj_iff fun_eq_iff o_apply) + apply (simp add: surj_iff_all) apply record_auto done @@ -386,7 +386,7 @@ done lemma surj_sysOfClient [iff]: "surj sysOfClient" - apply (simp add: surj_iff fun_eq_iff o_apply) + apply (simp add: surj_iff_all) apply record_auto done @@ -410,7 +410,7 @@ done lemma surj_client_map [iff]: "surj client_map" - apply (simp add: surj_iff fun_eq_iff o_apply) + apply (simp add: surj_iff_all) apply record_auto done @@ -682,7 +682,7 @@ lemma fst_lift_map_eq_fst [simp]: "fst (lift_map i x) i = fst x" apply (insert fst_o_lift_map [of i]) apply (drule fun_cong [where x=x]) -apply (simp add: o_def); +apply (simp add: o_def) done lemma fst_o_lift_map' [simp]: @@ -702,7 +702,7 @@ RS guarantees_PLam_I RS (bij_sysOfClient RS rename_rename_guarantees_eq RS iffD2) |> simplify (simpset() addsimps [lift_image_eq_rename, o_def, split_def, - surj_rename RS surj_range]) + surj_rename]) However, the "preserves" property remains to be discharged, and the unfolding of "o" and "sub" complicates subsequent reasoning. @@ -723,7 +723,7 @@ asm_simp_tac (ss addsimps [@{thm lift_guarantees_eq_lift_inv}, @{thm rename_guarantees_eq_rename_inv}, - @{thm bij_imp_bij_inv}, @{thm surj_rename} RS @{thm surj_range}, + @{thm bij_imp_bij_inv}, @{thm surj_rename}, @{thm inv_inv_eq}]) 1, asm_simp_tac (@{simpset} addsimps [@{thm o_def}, @{thm non_dummy_def}, @{thm guarantees_Int_right}]) 1] @@ -798,9 +798,9 @@ lemmas rename_Alloc_Increasing = Alloc_Increasing [THEN rename_guarantees_sysOfAlloc_I, - simplified surj_rename [THEN surj_range] o_def sub_apply + simplified surj_rename o_def sub_apply rename_image_Increasing bij_sysOfAlloc - allocGiv_o_inv_sysOfAlloc_eq']; + allocGiv_o_inv_sysOfAlloc_eq'] lemma System_Increasing_allocGiv: "i < Nclients ==> System : Increasing (sub i o allocGiv)" @@ -879,7 +879,7 @@ \ NbT + (\i \ lessThan Nclients. (tokens o sub i o allocRel) s)}" apply (simp add: o_apply) apply (insert Alloc_Safety [THEN rename_guarantees_sysOfAlloc_I]) - apply (simp add: o_def); + apply (simp add: o_def) apply (erule component_guaranteesD) apply (auto simp add: System_Increasing_allocRel [simplified sub_apply o_def]) done diff -r e7aa34600c36 -r 407c6122956f src/HOL/UNITY/Extend.thy --- a/src/HOL/UNITY/Extend.thy Fri Nov 26 09:15:49 2010 +0100 +++ b/src/HOL/UNITY/Extend.thy Fri Nov 26 10:04:04 2010 +0100 @@ -127,7 +127,7 @@ assumes surj_h: "surj h" and prem: "!! x y. g (h(x,y)) = x" shows "fst (inv h z) = g z" -by (metis UNIV_I f_inv_into_f pair_collapse prem surj_h surj_range) +by (metis UNIV_I f_inv_into_f pair_collapse prem surj_h) subsection{*Trivial properties of f, g, h*} diff -r e7aa34600c36 -r 407c6122956f src/HOL/UNITY/Rename.thy --- a/src/HOL/UNITY/Rename.thy Fri Nov 26 09:15:49 2010 +0100 +++ b/src/HOL/UNITY/Rename.thy Fri Nov 26 10:04:04 2010 +0100 @@ -60,7 +60,8 @@ lemma bij_extend_act: "bij h ==> bij (extend_act (%(x,u::'c). h x))" apply (rule bijI) apply (rule Extend.inj_extend_act) -apply (auto simp add: bij_extend_act_eq_project_act) +apply simp +apply (simp add: bij_extend_act_eq_project_act) apply (rule surjI) apply (rule Extend.extend_act_inverse) apply (blast intro: bij_imp_bij_inv good_map_bij) @@ -75,7 +76,7 @@ project_act (%(x,u::'c). h x)" apply (simp (no_asm_simp) add: bij_extend_act_eq_project_act [symmetric]) apply (rule surj_imp_inv_eq) -apply (blast intro: bij_extend_act bij_is_surj) +apply (blast intro!: bij_extend_act bij_is_surj) apply (simp (no_asm_simp) add: Extend.extend_act_inverse) done