src/HOL/Isar_Examples/Schroeder_Bernstein.thy
changeset 63980 f8e556c8ad6f
parent 63979 95c3ae4baba8
child 63981 6f7db4f8df4c
equal deleted inserted replaced
63979:95c3ae4baba8 63980:f8e556c8ad6f
     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:
       
    13   \<^item> \<^file>\<open>$ISABELLE_HOME/src/HOL/ex/Set_Theory.thy\<close>
       
    14   \<^item> \<^url>\<open>http://planetmath.org/proofofschroederbernsteintheoremusingtarskiknastertheorem\<close>
       
    15   \<^item> Springer LNCS 828 (cover page)
       
    16 \<close>
       
    17 
       
    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>
       
    20 proof
       
    21   define A where \<open>A = lfp (\<lambda>X. - (g ` (- (f ` X))))\<close>
       
    22   define g' where \<open>g' = inv g\<close>
       
    23   let \<open>?h\<close> = \<open>\<lambda>z. if z \<in> A then f z else g' z\<close>
       
    24 
       
    25   have \<open>A = - (g ` (- (f ` A)))\<close>
       
    26     unfolding A_def by (rule lfp_unfold) (blast intro: monoI)
       
    27   then have A_compl: \<open>- A = g ` (- (f ` A))\<close> by blast
       
    28   then have *: \<open>g' ` (- A) = - (f ` A)\<close>
       
    29     using g'_def \<open>inj g\<close> by auto
       
    30 
       
    31   show \<open>inj ?h \<and> surj ?h\<close>
       
    32   proof
       
    33     from * show \<open>surj ?h\<close> by auto
       
    34     have \<open>inj_on f A\<close>
       
    35       using \<open>inj f\<close> by (rule subset_inj_on) blast
       
    36     moreover
       
    37     have \<open>inj_on g' (- A)\<close>
       
    38       unfolding g'_def
       
    39     proof (rule inj_on_inv_into)
       
    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)
       
    42     qed
       
    43     moreover
       
    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
       
    45     proof -
       
    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
       
    49       with eq fa show \<open>False\<close> by simp
       
    50     qed
       
    51     ultimately show \<open>inj ?h\<close>
       
    52       unfolding inj_on_def by (metis ComplI)
       
    53   qed
       
    54 qed
       
    55 
       
    56 end