src/HOL/Isar_Examples/Schroeder_Bernstein.thy
changeset 63297 ce995deef4b0
parent 63291 b1d7950285cf
child 63583 a39baba12732
equal deleted inserted replaced
63296:3951a15a05d1 63297:ce995deef4b0
    13   \<^item> @{file "$ISABELLE_HOME/src/HOL/ex/Set_Theory.thy"}
    13   \<^item> @{file "$ISABELLE_HOME/src/HOL/ex/Set_Theory.thy"}
    14   \<^item> @{url "http://planetmath.org/proofofschroederbernsteintheoremusingtarskiknastertheorem"}
    14   \<^item> @{url "http://planetmath.org/proofofschroederbernsteintheoremusingtarskiknastertheorem"}
    15   \<^item> Springer LNCS 828 (cover page)
    15   \<^item> Springer LNCS 828 (cover page)
    16 \<close>
    16 \<close>
    17 
    17 
    18 theorem Schroeder_Bernstein:
    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   fixes f :: \<open>'a \<Rightarrow> 'b\<close>
    19   for f :: \<open>'a \<Rightarrow> 'b\<close> and g :: \<open>'b \<Rightarrow> 'a\<close>
    20     and g :: \<open>'b \<Rightarrow> 'a\<close>
       
    21   assumes \<open>inj f\<close> and \<open>inj g\<close>
       
    22   shows \<open>\<exists>h :: 'a \<Rightarrow> 'b. inj h \<and> surj h\<close>
       
    23 proof
    20 proof
    24   define A where \<open>A = lfp (\<lambda>X. - (g ` (- (f ` X))))\<close>
    21   define A where \<open>A = lfp (\<lambda>X. - (g ` (- (f ` X))))\<close>
    25   define g' where \<open>g' = inv g\<close>
    22   define g' where \<open>g' = inv g\<close>
    26   let ?h = \<open>\<lambda>z. if z \<in> A then f z else g' z\<close>
    23   let \<open>?h\<close> = \<open>\<lambda>z. if z \<in> A then f z else g' z\<close>
    27 
    24 
    28   have \<open>A = - (g ` (- (f ` A)))\<close>
    25   have \<open>A = - (g ` (- (f ` A)))\<close>
    29     unfolding A_def by (rule lfp_unfold) (blast intro: monoI)
    26     unfolding A_def by (rule lfp_unfold) (blast intro: monoI)
    30   then have A_compl: \<open>- A = g ` (- (f ` A))\<close> by blast
    27   then have A_compl: \<open>- A = g ` (- (f ` A))\<close> by blast
    31   then have *: \<open>g' ` (- A) = - (f ` A)\<close>
    28   then have *: \<open>g' ` (- A) = - (f ` A)\<close>
    42     proof (rule inj_on_inv_into)
    39     proof (rule inj_on_inv_into)
    43       have \<open>g ` (- (f ` A)) \<subseteq> range g\<close> by blast
    40       have \<open>g ` (- (f ` A)) \<subseteq> range g\<close> by blast
    44       then show \<open>- A \<subseteq> range g\<close> by (simp only: A_compl)
    41       then show \<open>- A \<subseteq> range g\<close> by (simp only: A_compl)
    45     qed
    42     qed
    46     moreover
    43     moreover
    47     have False 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
    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
    48     proof -
    45     proof -
    49       from a have fa: \<open>f a \<in> f ` A\<close> by (rule imageI)
    46       from a have fa: \<open>f a \<in> f ` A\<close> by (rule imageI)
    50       from b have \<open>g' b \<in> g' ` (- A)\<close> by (rule imageI)
    47       from b have \<open>g' b \<in> g' ` (- A)\<close> by (rule imageI)
    51       with * have \<open>g' b \<in> - (f ` A)\<close> by simp
    48       with * have \<open>g' b \<in> - (f ` A)\<close> by simp
    52       with eq fa show False by simp
    49       with eq fa show \<open>False\<close> by simp
    53     qed
    50     qed
    54     ultimately show \<open>inj ?h\<close>
    51     ultimately show \<open>inj ?h\<close>
    55       unfolding inj_on_def by (metis ComplI)
    52       unfolding inj_on_def by (metis ComplI)
    56   qed
    53   qed
    57 qed
    54 qed