# HG changeset patch # User wenzelm # Date 1465850558 -7200 # Node ID ce995deef4b07e747da7a77e5bed096ee13760d3 # Parent 3951a15a05d1285f423ee10399d3828035f09fa6 tuned; diff -r 3951a15a05d1 -r ce995deef4b0 src/HOL/Isar_Examples/Schroeder_Bernstein.thy --- a/src/HOL/Isar_Examples/Schroeder_Bernstein.thy Mon Jun 13 17:39:52 2016 +0200 +++ b/src/HOL/Isar_Examples/Schroeder_Bernstein.thy Mon Jun 13 22:42:38 2016 +0200 @@ -15,15 +15,12 @@ \<^item> Springer LNCS 828 (cover page) \ -theorem Schroeder_Bernstein: - fixes f :: \'a \ 'b\ - and g :: \'b \ 'a\ - assumes \inj f\ and \inj g\ - shows \\h :: 'a \ 'b. inj h \ surj h\ +theorem Schroeder_Bernstein: \\h :: 'a \ 'b. inj h \ surj h\ if \inj f\ and \inj g\ + for f :: \'a \ 'b\ and g :: \'b \ 'a\ proof define A where \A = lfp (\X. - (g ` (- (f ` X))))\ define g' where \g' = inv g\ - let ?h = \\z. if z \ A then f z else g' z\ + let \?h\ = \\z. if z \ A then f z else g' z\ have \A = - (g ` (- (f ` A)))\ unfolding A_def by (rule lfp_unfold) (blast intro: monoI) @@ -44,12 +41,12 @@ then show \- A \ range g\ by (simp only: A_compl) qed moreover - have False if eq: \f a = g' b\ and a: \a \ A\ and b: \b \ - A\ for a b + have \False\ if eq: \f a = g' b\ and a: \a \ A\ and b: \b \ - A\ for a b proof - from a have fa: \f a \ f ` A\ by (rule imageI) from b have \g' b \ g' ` (- A)\ by (rule imageI) with * have \g' b \ - (f ` A)\ by simp - with eq fa show False by simp + with eq fa show \False\ by simp qed ultimately show \inj ?h\ unfolding inj_on_def by (metis ComplI)