| 61938 |      1 | (*  Title:      HOL/Isar_Examples/Schroeder_Bernstein.thy
 | 
|  |      2 |     Author:     Makarius
 | 
|  |      3 | *)
 | 
|  |      4 | 
 | 
|  |      5 | section \<open>Schröder-Bernstein Theorem\<close>
 | 
|  |      6 | 
 | 
|  |      7 | theory Schroeder_Bernstein
 | 
|  |      8 | imports Main
 | 
|  |      9 | begin
 | 
|  |     10 | 
 | 
|  |     11 | text \<open>
 | 
|  |     12 |   See also:
 | 
| 61939 |     13 |   \<^item> @{file "$ISABELLE_HOME/src/HOL/ex/Set_Theory.thy"}
 | 
| 61938 |     14 |   \<^item> @{url "http://planetmath.org/proofofschroederbernsteintheoremusingtarskiknastertheorem"}
 | 
|  |     15 |   \<^item> Springer LNCS 828 (cover page)
 | 
|  |     16 | \<close>
 | 
|  |     17 | 
 | 
| 63297 |     18 | theorem Schroeder_Bernstein: \<open>\<exists>h :: 'a \<Rightarrow> 'b. inj h \<and> surj h\<close> if \<open>inj f\<close> and \<open>inj g\<close>
 | 
|  |     19 |   for f :: \<open>'a \<Rightarrow> 'b\<close> and g :: \<open>'b \<Rightarrow> 'a\<close>
 | 
| 61938 |     20 | proof
 | 
| 63291 |     21 |   define A where \<open>A = lfp (\<lambda>X. - (g ` (- (f ` X))))\<close>
 | 
|  |     22 |   define g' where \<open>g' = inv g\<close>
 | 
| 63297 |     23 |   let \<open>?h\<close> = \<open>\<lambda>z. if z \<in> A then f z else g' z\<close>
 | 
| 61938 |     24 | 
 | 
| 63291 |     25 |   have \<open>A = - (g ` (- (f ` A)))\<close>
 | 
| 61938 |     26 |     unfolding A_def by (rule lfp_unfold) (blast intro: monoI)
 | 
| 63291 |     27 |   then have A_compl: \<open>- A = g ` (- (f ` A))\<close> by blast
 | 
|  |     28 |   then have *: \<open>g' ` (- A) = - (f ` A)\<close>
 | 
| 61938 |     29 |     using g'_def \<open>inj g\<close> by auto
 | 
|  |     30 | 
 | 
| 63291 |     31 |   show \<open>inj ?h \<and> surj ?h\<close>
 | 
| 61938 |     32 |   proof
 | 
| 63291 |     33 |     from * show \<open>surj ?h\<close> by auto
 | 
|  |     34 |     have \<open>inj_on f A\<close>
 | 
| 61938 |     35 |       using \<open>inj f\<close> by (rule subset_inj_on) blast
 | 
|  |     36 |     moreover
 | 
| 63291 |     37 |     have \<open>inj_on g' (- A)\<close>
 | 
| 61938 |     38 |       unfolding g'_def
 | 
|  |     39 |     proof (rule inj_on_inv_into)
 | 
| 63291 |     40 |       have \<open>g ` (- (f ` A)) \<subseteq> range g\<close> by blast
 | 
|  |     41 |       then show \<open>- A \<subseteq> range g\<close> by (simp only: A_compl)
 | 
| 61938 |     42 |     qed
 | 
|  |     43 |     moreover
 | 
| 63297 |     44 |     have \<open>False\<close> if eq: \<open>f a = g' b\<close> and a: \<open>a \<in> A\<close> and b: \<open>b \<in> - A\<close> for a b
 | 
| 61938 |     45 |     proof -
 | 
| 63291 |     46 |       from a have fa: \<open>f a \<in> f ` A\<close> by (rule imageI)
 | 
|  |     47 |       from b have \<open>g' b \<in> g' ` (- A)\<close> by (rule imageI)
 | 
|  |     48 |       with * have \<open>g' b \<in> - (f ` A)\<close> by simp
 | 
| 63297 |     49 |       with eq fa show \<open>False\<close> by simp
 | 
| 61938 |     50 |     qed
 | 
| 63291 |     51 |     ultimately show \<open>inj ?h\<close>
 | 
| 61938 |     52 |       unfolding inj_on_def by (metis ComplI)
 | 
|  |     53 |   qed
 | 
|  |     54 | qed
 | 
|  |     55 | 
 | 
|  |     56 | end
 |