merged
authorwenzelm
Sat, 01 Oct 2016 20:03:17 +0200
changeset 63983 db9259a5e20e
parent 63978 efc958d2fe00 (current diff)
parent 63982 4c4049e3bad8 (diff)
child 63984 6ba87450894d
merged
src/HOL/Isar_Examples/Schroeder_Bernstein.thy
--- a/NEWS	Sat Oct 01 15:21:43 2016 +0200
+++ b/NEWS	Sat Oct 01 20:03:17 2016 +0200
@@ -336,6 +336,12 @@
 eliminates the need to qualify any of those names in the presence of
 infix "mod" syntax. INCOMPATIBILITY.
 
+* Statements and proofs of Knaster-Tarski fixpoint combinators lfp/gfp
+have been clarified. The fixpoint properties are lfp_fixpoint, its
+symmetric lfp_unfold (as before), and the duals for gfp. Auxiliary items
+for the proof (lfp_lemma2 etc.) are no longer exported, but can be
+easily recovered by composition with eq_refl. Minor INCOMPATIBILITY.
+
 * Constant "surj" is a mere input abbreviation, to avoid hiding an
 equation in term output. Minor INCOMPATIBILITY.
 
--- a/src/HOL/BNF_Cardinal_Order_Relation.thy	Sat Oct 01 15:21:43 2016 +0200
+++ b/src/HOL/BNF_Cardinal_Order_Relation.thy	Sat Oct 01 20:03:17 2016 +0200
@@ -183,7 +183,7 @@
    card_of_Well_order[of "B"] Field_card_of[of "B"] Field_card_of[of "A"]
    embed_Field[of "|B|" "|A|" g] by auto
    obtain h where "bij_betw h A B"
-   using * ** 1 Cantor_Bernstein[of f] by fastforce
+   using * ** 1 Schroeder_Bernstein[of f] by fastforce
    hence "|A| =o |B|" using card_of_ordIso by blast
    hence "|A| \<le>o |B|" using ordIso_iff_ordLeq by auto
   }
--- a/src/HOL/Complete_Partial_Order.thy	Sat Oct 01 15:21:43 2016 +0200
+++ b/src/HOL/Complete_Partial_Order.thy	Sat Oct 01 20:03:17 2016 +0200
@@ -365,15 +365,15 @@
   by standard (fast intro: Sup_upper Sup_least)+
 
 lemma lfp_eq_fixp:
-  assumes f: "mono f"
+  assumes mono: "mono f"
   shows "lfp f = fixp f"
 proof (rule antisym)
-  from f have f': "monotone (op \<le>) (op \<le>) f"
+  from mono have f': "monotone (op \<le>) (op \<le>) f"
     unfolding mono_def monotone_def .
   show "lfp f \<le> fixp f"
     by (rule lfp_lowerbound, subst fixp_unfold [OF f'], rule order_refl)
   show "fixp f \<le> lfp f"
-    by (rule fixp_lowerbound [OF f'], subst lfp_unfold [OF f], rule order_refl)
+    by (rule fixp_lowerbound [OF f']) (simp add: lfp_fixpoint [OF mono])
 qed
 
 hide_const (open) iterates fixp
--- a/src/HOL/Finite_Set.thy	Sat Oct 01 15:21:43 2016 +0200
+++ b/src/HOL/Finite_Set.thy	Sat Oct 01 20:03:17 2016 +0200
@@ -139,12 +139,12 @@
   qed
 qed
 
-lemma finite_conv_nat_seg_image: "finite A \<longleftrightarrow> (\<exists>(n::nat) f. A = f ` {i::nat. i < n})"
+lemma finite_conv_nat_seg_image: "finite A \<longleftrightarrow> (\<exists>n f. A = f ` {i::nat. i < n})"
   by (blast intro: nat_seg_image_imp_finite dest: finite_imp_nat_seg_image_inj_on)
 
 lemma finite_imp_inj_to_nat_seg:
   assumes "finite A"
-  shows "\<exists>f n::nat. f ` A = {i. i < n} \<and> inj_on f A"
+  shows "\<exists>f n. f ` A = {i::nat. i < n} \<and> inj_on f A"
 proof -
   from finite_imp_nat_seg_image_inj_on [OF \<open>finite A\<close>]
   obtain f and n :: nat where bij: "bij_betw f {i. i<n} A"
--- a/src/HOL/Hilbert_Choice.thy	Sat Oct 01 15:21:43 2016 +0200
+++ b/src/HOL/Hilbert_Choice.thy	Sat Oct 01 20:03:17 2016 +0200
@@ -450,106 +450,6 @@
   using fg gf inv_equality[of g f] by (auto simp add: fun_eq_iff)
 
 
-subsection \<open>The Cantor-Bernstein Theorem\<close>
-
-lemma Cantor_Bernstein_aux:
-  "\<exists>A' h. A' \<subseteq> A \<and>
-    (\<forall>a \<in> A'. a \<notin> g ` (B - f ` A')) \<and>
-    (\<forall>a \<in> A'. h a = f a) \<and>
-    (\<forall>a \<in> A - A'. h a \<in> B - (f ` A') \<and> a = g (h a))"
-proof -
-  define H where "H A' = A - (g ` (B - (f ` A')))" for A'
-  have "mono H" unfolding mono_def H_def by blast
-  from lfp_unfold [OF this] obtain A' where "H A' = A'" by blast
-  then have "A' = A - (g ` (B - (f ` A')))" by (simp add: H_def)
-  then have 1: "A' \<subseteq> A"
-    and 2: "\<forall>a \<in> A'.  a \<notin> g ` (B - f ` A')"
-    and 3: "\<forall>a \<in> A - A'. \<exists>b \<in> B - (f ` A'). a = g b"
-    by blast+
-  define h where "h a = (if a \<in> A' then f a else (SOME b. b \<in> B - (f ` A') \<and> a = g b))" for a
-  then have 4: "\<forall>a \<in> A'. h a = f a" by simp
-  have "\<forall>a \<in> A - A'. h a \<in> B - (f ` A') \<and> a = g (h a)"
-  proof
-    fix a
-    let ?phi = "\<lambda>b. b \<in> B - (f ` A') \<and> a = g b"
-    assume *: "a \<in> A - A'"
-    from * have "h a = (SOME b. ?phi b)" by (auto simp: h_def)
-    moreover from 3 * have "\<exists>b. ?phi b" by auto
-    ultimately show "?phi (h a)"
-      using someI_ex[of ?phi] by auto
-  qed
-  with 1 2 4 show ?thesis by blast
-qed
-
-theorem Cantor_Bernstein:
-  assumes inj1: "inj_on f A" and sub1: "f ` A \<subseteq> B"
-    and inj2: "inj_on g B" and sub2: "g ` B \<subseteq> A"
-  shows "\<exists>h. bij_betw h A B"
-proof-
-  obtain A' and h where "A' \<subseteq> A"
-    and 1: "\<forall>a \<in> A'. a \<notin> g ` (B - f ` A')"
-    and 2: "\<forall>a \<in> A'. h a = f a"
-    and 3: "\<forall>a \<in> A - A'. h a \<in> B - (f ` A') \<and> 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 \<in> A" and 5: "a2 \<in> A" and 6: "h a1 = h a2"
-    show "a1 = a2"
-    proof (cases "a1 \<in> A'")
-      case True
-      show ?thesis
-      proof (cases "a2 \<in> A'")
-        case True': True
-        with True 2 6 have "f a1 = f a2" by auto
-        with inj1 \<open>A' \<subseteq> A\<close> True True' show ?thesis
-          unfolding inj_on_def by blast
-      next
-        case False
-        with 2 3 5 6 True have False by force
-        then show ?thesis ..
-      qed
-    next
-      case False
-      show ?thesis
-      proof (cases "a2 \<in> A'")
-        case True
-        with 2 3 4 6 False have False by auto
-        then show ?thesis ..
-      next
-        case False': False
-        with False 3 4 5 have "a1 = g (h a1)" "a2 = g (h a2)" by auto
-        with 6 show ?thesis by simp
-      qed
-    qed
-  qed
-  moreover have "h ` A = B"
-  proof safe
-    fix a
-    assume "a \<in> A"
-    with sub1 2 3 show "h a \<in> B" by (cases "a \<in> A'") auto
-  next
-    fix b
-    assume *: "b \<in> B"
-    show "b \<in> h ` A"
-    proof (cases "b \<in> f ` A'")
-      case True
-      then obtain a where "a \<in> A'" "b = f a" by blast
-      with \<open>A' \<subseteq> A\<close> 2 show ?thesis by force
-    next
-      case False
-      with 1 * have "g b \<notin> A'" by auto
-      with sub2 * have 4: "g b \<in> A - A'" by auto
-      with 3 have "h (g b) \<in> B" "g (h (g b)) = g b" by auto
-      with inj2 * have "h (g b) = b" by (auto simp: inj_on_def)
-      with 4 show ?thesis by force
-    qed
-  qed
-  ultimately show ?thesis
-    by (auto simp: bij_betw_def)
-qed
-
-
 subsection \<open>Other Consequences of Hilbert's Epsilon\<close>
 
 text \<open>Hilbert's Epsilon and the @{term split} Operator\<close>
@@ -566,30 +466,48 @@
 
 
 text \<open>A relation is wellfounded iff it has no infinite descending chain.\<close>
-lemma wf_iff_no_infinite_down_chain: "wf r \<longleftrightarrow> (\<not> (\<exists>f. \<forall>i. (f (Suc i), f i) \<in> r))"
-  apply (simp only: wf_eq_minimal)
-  apply (rule iffI)
-   apply (rule notI)
-   apply (erule exE)
-   apply (erule_tac x = "{w. \<exists>i. w = f i}" in allE)
-   apply blast
-  apply (erule contrapos_np)
-  apply simp
-  apply clarify
-  apply (subgoal_tac "\<forall>n. rec_nat x (\<lambda>i y. SOME z. z \<in> Q \<and> (z, y) \<in> r) n \<in> Q")
-   apply (rule_tac x = "rec_nat x (\<lambda>i y. SOME z. z \<in> Q \<and> (z, y) \<in> r)" in exI)
-   apply (rule allI)
-   apply simp
-   apply (rule someI2_ex)
-    apply blast
-   apply blast
-  apply (rule allI)
-  apply (induct_tac n)
-   apply simp_all
-  apply (rule someI2_ex)
-   apply blast
-  apply blast
-  done
+lemma wf_iff_no_infinite_down_chain: "wf r \<longleftrightarrow> (\<nexists>f. \<forall>i. (f (Suc i), f i) \<in> r)"
+  (is "_ \<longleftrightarrow> \<not> ?ex")
+proof
+  assume "wf r"
+  show "\<not> ?ex"
+  proof
+    assume ?ex
+    then obtain f where f: "(f (Suc i), f i) \<in> r" for i
+      by blast
+    from \<open>wf r\<close> have minimal: "x \<in> Q \<Longrightarrow> \<exists>z\<in>Q. \<forall>y. (y, z) \<in> r \<longrightarrow> y \<notin> Q" for x Q
+      by (auto simp: wf_eq_minimal)
+    let ?Q = "{w. \<exists>i. w = f i}"
+    fix n
+    have "f n \<in> ?Q" by blast
+    from minimal [OF this] obtain j where "(y, f j) \<in> r \<Longrightarrow> y \<notin> ?Q" for y by blast
+    with this [OF \<open>(f (Suc j), f j) \<in> r\<close>] have "f (Suc j) \<notin> ?Q" by simp
+    then show False by blast
+  qed
+next
+  assume "\<not> ?ex"
+  then show "wf r"
+  proof (rule contrapos_np)
+    assume "\<not> wf r"
+    then obtain Q x where x: "x \<in> Q" and rec: "z \<in> Q \<Longrightarrow> \<exists>y. (y, z) \<in> r \<and> y \<in> Q" for z
+      by (auto simp add: wf_eq_minimal)
+    obtain descend :: "nat \<Rightarrow> 'a"
+      where descend_0: "descend 0 = x"
+        and descend_Suc: "descend (Suc n) = (SOME y. y \<in> Q \<and> (y, descend n) \<in> r)" for n
+      by (rule that [of "rec_nat x (\<lambda>_ rec. (SOME y. y \<in> Q \<and> (y, rec) \<in> r))"]) simp_all
+    have descend_Q: "descend n \<in> Q" for n
+    proof (induct n)
+      case 0
+      with x show ?case by (simp only: descend_0)
+    next
+      case Suc
+      then show ?case by (simp only: descend_Suc) (rule someI2_ex; use rec in blast)
+    qed
+    have "(descend (Suc i), descend i) \<in> r" for i
+      by (simp only: descend_Suc) (rule someI2_ex; use descend_Q rec in blast)
+    then show "\<exists>f. \<forall>i. (f (Suc i), f i) \<in> r" by blast
+  qed
+qed
 
 lemma wf_no_infinite_down_chainE:
   assumes "wf r"
--- a/src/HOL/Inductive.thy	Sat Oct 01 15:21:43 2016 +0200
+++ b/src/HOL/Inductive.thy	Sat Oct 01 20:03:17 2016 +0200
@@ -14,38 +14,47 @@
     "primrec" :: thy_decl
 begin
 
-subsection \<open>Least and greatest fixed points\<close>
+subsection \<open>Least fixed points\<close>
 
 context complete_lattice
 begin
 
-definition lfp :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a"  \<comment> \<open>least fixed point\<close>
+definition lfp :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a"
   where "lfp f = Inf {u. f u \<le> u}"
 
-definition gfp :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a"  \<comment> \<open>greatest fixed point\<close>
-  where "gfp f = Sup {u. u \<le> f u}"
-
-
-subsection \<open>Proof of Knaster-Tarski Theorem using @{term lfp}\<close>
-
-text \<open>@{term "lfp f"} is the least upper bound of the set @{term "{u. f u \<le> u}"}\<close>
-
 lemma lfp_lowerbound: "f A \<le> A \<Longrightarrow> lfp f \<le> A"
-  by (auto simp add: lfp_def intro: Inf_lower)
+  unfolding lfp_def by (rule Inf_lower) simp
 
 lemma lfp_greatest: "(\<And>u. f u \<le> u \<Longrightarrow> A \<le> u) \<Longrightarrow> A \<le> lfp f"
-  by (auto simp add: lfp_def intro: Inf_greatest)
+  unfolding lfp_def by (rule Inf_greatest) simp
 
 end
 
-lemma lfp_lemma2: "mono f \<Longrightarrow> f (lfp f) \<le> lfp f"
-  by (iprover intro: lfp_greatest order_trans monoD lfp_lowerbound)
-
-lemma lfp_lemma3: "mono f \<Longrightarrow> lfp f \<le> f (lfp f)"
-  by (iprover intro: lfp_lemma2 monoD lfp_lowerbound)
+lemma lfp_fixpoint:
+  assumes "mono f"
+  shows "f (lfp f) = lfp f"
+  unfolding lfp_def
+proof (rule order_antisym)
+  let ?H = "{u. f u \<le> u}"
+  let ?a = "\<Sqinter>?H"
+  show "f ?a \<le> ?a"
+  proof (rule Inf_greatest)
+    fix x
+    assume "x \<in> ?H"
+    then have "?a \<le> x" by (rule Inf_lower)
+    with \<open>mono f\<close> have "f ?a \<le> f x" ..
+    also from \<open>x \<in> ?H\<close> have "f x \<le> x" ..
+    finally show "f ?a \<le> x" .
+  qed
+  show "?a \<le> f ?a"
+  proof (rule Inf_lower)
+    from \<open>mono f\<close> and \<open>f ?a \<le> ?a\<close> have "f (f ?a) \<le> f ?a" ..
+    then show "f ?a \<in> ?H" ..
+  qed
+qed
 
 lemma lfp_unfold: "mono f \<Longrightarrow> lfp f = f (lfp f)"
-  by (iprover intro: order_antisym lfp_lemma2 lfp_lemma3)
+  by (rule lfp_fixpoint [symmetric])
 
 lemma lfp_const: "lfp (\<lambda>x. t) = t"
   by (rule lfp_unfold) (simp add: mono_def)
@@ -132,9 +141,13 @@
   by (rule lfp_lowerbound [THEN lfp_greatest]) (blast intro: order_trans)
 
 
-subsection \<open>Proof of Knaster-Tarski Theorem using \<open>gfp\<close>\<close>
+subsection \<open>Greatest fixed points\<close>
 
-text \<open>@{term "gfp f"} is the greatest lower bound of the set @{term "{u. u \<le> f u}"}\<close>
+context complete_lattice
+begin
+
+definition gfp :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a"
+  where "gfp f = Sup {u. u \<le> f u}"
 
 lemma gfp_upperbound: "X \<le> f X \<Longrightarrow> X \<le> gfp f"
   by (auto simp add: gfp_def intro: Sup_upper)
@@ -142,14 +155,36 @@
 lemma gfp_least: "(\<And>u. u \<le> f u \<Longrightarrow> u \<le> X) \<Longrightarrow> gfp f \<le> X"
   by (auto simp add: gfp_def intro: Sup_least)
 
-lemma gfp_lemma2: "mono f \<Longrightarrow> gfp f \<le> f (gfp f)"
-  by (iprover intro: gfp_least order_trans monoD gfp_upperbound)
+end
+
+lemma lfp_le_gfp: "mono f \<Longrightarrow> lfp f \<le> gfp f"
+  by (rule gfp_upperbound) (simp add: lfp_fixpoint)
 
-lemma gfp_lemma3: "mono f \<Longrightarrow> f (gfp f) \<le> gfp f"
-  by (iprover intro: gfp_lemma2 monoD gfp_upperbound)
+lemma gfp_fixpoint:
+  assumes "mono f"
+  shows "f (gfp f) = gfp f"
+  unfolding gfp_def
+proof (rule order_antisym)
+  let ?H = "{u. u \<le> f u}"
+  let ?a = "\<Squnion>?H"
+  show "?a \<le> f ?a"
+  proof (rule Sup_least)
+    fix x
+    assume "x \<in> ?H"
+    then have "x \<le> f x" ..
+    also from \<open>x \<in> ?H\<close> have "x \<le> ?a" by (rule Sup_upper)
+    with \<open>mono f\<close> have "f x \<le> f ?a" ..
+    finally show "x \<le> f ?a" .
+  qed
+  show "f ?a \<le> ?a"
+  proof (rule Sup_upper)
+    from \<open>mono f\<close> and \<open>?a \<le> f ?a\<close> have "f ?a \<le> f (f ?a)" ..
+    then show "f ?a \<in> ?H" ..
+  qed
+qed
 
 lemma gfp_unfold: "mono f \<Longrightarrow> gfp f = f (gfp f)"
-  by (iprover intro: order_antisym gfp_lemma2 gfp_lemma3)
+  by (rule gfp_fixpoint [symmetric])
 
 lemma gfp_const: "gfp (\<lambda>x. t) = t"
   by (rule gfp_unfold) (simp add: mono_def)
@@ -158,10 +193,6 @@
   by (rule antisym) (simp_all add: gfp_upperbound gfp_unfold[symmetric])
 
 
-lemma lfp_le_gfp: "mono f \<Longrightarrow> lfp f \<le> gfp f"
-  by (iprover intro: gfp_upperbound lfp_lemma3)
-
-
 subsection \<open>Coinduction rules for greatest fixed points\<close>
 
 text \<open>Weak version.\<close>
@@ -174,7 +205,7 @@
   done
 
 lemma coinduct_lemma: "X \<le> f (sup X (gfp f)) \<Longrightarrow> mono f \<Longrightarrow> sup X (gfp f) \<le> f (sup X (gfp f))"
-  apply (frule gfp_lemma2)
+  apply (frule gfp_unfold [THEN eq_refl])
   apply (drule mono_sup)
   apply (rule le_supI)
    apply assumption
@@ -190,7 +221,7 @@
   by (rule weak_coinduct[rotated], rule coinduct_lemma) blast+
 
 lemma gfp_fun_UnI2: "mono f \<Longrightarrow> a \<in> gfp f \<Longrightarrow> a \<in> f (X \<union> gfp f)"
-  by (blast dest: gfp_lemma2 mono_Un)
+  by (blast dest: gfp_fixpoint mono_Un)
 
 lemma gfp_ordinal_induct[case_names mono step union]:
   fixes f :: "'a::complete_lattice \<Rightarrow> 'a"
@@ -248,7 +279,7 @@
   "X \<subseteq> f (lfp (\<lambda>x. f x \<union> X \<union> gfp f)) \<Longrightarrow> mono f \<Longrightarrow>
     lfp (\<lambda>x. f x \<union> X \<union> gfp f) \<subseteq> f (lfp (\<lambda>x. f x \<union> X \<union> gfp f))"
   apply (rule subset_trans)
-   apply (erule coinduct3_mono_lemma [THEN lfp_lemma3])
+   apply (erule coinduct3_mono_lemma [THEN lfp_unfold [THEN eq_refl]])
   apply (rule Un_least [THEN Un_least])
     apply (rule subset_refl, assumption)
   apply (rule gfp_unfold [THEN equalityD1, THEN subset_trans], assumption)
@@ -391,6 +422,94 @@
   induct_rulify_fallback
 
 
+subsection \<open>The Schroeder-Bernstein Theorem\<close>
+
+text \<open>
+  See also:
+  \<^item> \<^file>\<open>$ISABELLE_HOME/src/HOL/ex/Set_Theory.thy\<close>
+  \<^item> \<^url>\<open>http://planetmath.org/proofofschroederbernsteintheoremusingtarskiknastertheorem\<close>
+  \<^item> Springer LNCS 828 (cover page)
+\<close>
+
+theorem Schroeder_Bernstein:
+  fixes f :: "'a \<Rightarrow> 'b" and g :: "'b \<Rightarrow> 'a"
+    and A :: "'a set" and B :: "'b set"
+  assumes inj1: "inj_on f A" and sub1: "f ` A \<subseteq> B"
+    and inj2: "inj_on g B" and sub2: "g ` B \<subseteq> A"
+  shows "\<exists>h. bij_betw h A B"
+proof (rule exI, rule bij_betw_imageI)
+  define X where "X = lfp (\<lambda>X. A - (g ` (B - (f ` X))))"
+  define g' where "g' = the_inv_into (B - (f ` X)) g"
+  let ?h = "\<lambda>z. if z \<in> X then f z else g' z"
+
+  have X: "X = A - (g ` (B - (f ` X)))"
+    unfolding X_def by (rule lfp_unfold) (blast intro: monoI)
+  then have X_compl: "A - X = g ` (B - (f ` X))"
+    using sub2 by blast
+
+  from inj2 have inj2': "inj_on g (B - (f ` X))"
+    by (rule inj_on_subset) auto
+  with X_compl have *: "g' ` (A - X) = B - (f ` X)"
+    by (simp add: g'_def)
+
+  from X have X_sub: "X \<subseteq> A" by auto
+  from X sub1 have fX_sub: "f ` X \<subseteq> B" by auto
+
+  show "?h ` A = B"
+  proof -
+    from X_sub have "?h ` A = ?h ` (X \<union> (A - X))" by auto
+    also have "\<dots> = ?h ` X \<union> ?h ` (A - X)" by (simp only: image_Un)
+    also have "?h ` X = f ` X" by auto
+    also from * have "?h ` (A - X) = B - (f ` X)" by auto
+    also from fX_sub have "f ` X \<union> (B - f ` X) = B" by blast
+    finally show ?thesis .
+  qed
+  show "inj_on ?h A"
+  proof -
+    from inj1 X_sub have on_X: "inj_on f X"
+      by (rule subset_inj_on)
+
+    have on_X_compl: "inj_on g' (A - X)"
+      unfolding g'_def X_compl
+      by (rule inj_on_the_inv_into) (rule inj2')
+
+    have impossible: False if eq: "f a = g' b" and a: "a \<in> X" and b: "b \<in> A - X" for a b
+    proof -
+      from a have fa: "f a \<in> f ` X" by (rule imageI)
+      from b have "g' b \<in> g' ` (A - X)" by (rule imageI)
+      with * have "g' b \<in> - (f ` X)" by simp
+      with eq fa show False by simp
+    qed
+
+    show ?thesis
+    proof (rule inj_onI)
+      fix a b
+      assume h: "?h a = ?h b"
+      assume "a \<in> A" and "b \<in> A"
+      then consider "a \<in> X" "b \<in> X" | "a \<in> A - X" "b \<in> A - X"
+        | "a \<in> X" "b \<in> A - X" | "a \<in> A - X" "b \<in> X"
+        by blast
+      then show "a = b"
+      proof cases
+        case 1
+        with h on_X show ?thesis by (simp add: inj_on_eq_iff)
+      next
+        case 2
+        with h on_X_compl show ?thesis by (simp add: inj_on_eq_iff)
+      next
+        case 3
+        with h impossible [of a b] have False by simp
+        then show ?thesis ..
+      next
+        case 4
+        with h impossible [of b a] have False by simp
+        then show ?thesis ..
+      qed
+    qed
+  qed
+qed
+
+
 subsection \<open>Inductive datatypes and primitive recursion\<close>
 
 text \<open>Package setup.\<close>
--- a/src/HOL/Isar_Examples/Schroeder_Bernstein.thy	Sat Oct 01 15:21:43 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,56 +0,0 @@
-(*  Title:      HOL/Isar_Examples/Schroeder_Bernstein.thy
-    Author:     Makarius
-*)
-
-section \<open>Schröder-Bernstein Theorem\<close>
-
-theory Schroeder_Bernstein
-  imports Main
-begin
-
-text \<open>
-  See also:
-  \<^item> \<^file>\<open>$ISABELLE_HOME/src/HOL/ex/Set_Theory.thy\<close>
-  \<^item> \<^url>\<open>http://planetmath.org/proofofschroederbernsteintheoremusingtarskiknastertheorem\<close>
-  \<^item> Springer LNCS 828 (cover page)
-\<close>
-
-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>
-  for f :: \<open>'a \<Rightarrow> 'b\<close> and g :: \<open>'b \<Rightarrow> 'a\<close>
-proof
-  define A where \<open>A = lfp (\<lambda>X. - (g ` (- (f ` X))))\<close>
-  define g' where \<open>g' = inv g\<close>
-  let \<open>?h\<close> = \<open>\<lambda>z. if z \<in> A then f z else g' z\<close>
-
-  have \<open>A = - (g ` (- (f ` A)))\<close>
-    unfolding A_def by (rule lfp_unfold) (blast intro: monoI)
-  then have A_compl: \<open>- A = g ` (- (f ` A))\<close> by blast
-  then have *: \<open>g' ` (- A) = - (f ` A)\<close>
-    using g'_def \<open>inj g\<close> by auto
-
-  show \<open>inj ?h \<and> surj ?h\<close>
-  proof
-    from * show \<open>surj ?h\<close> by auto
-    have \<open>inj_on f A\<close>
-      using \<open>inj f\<close> by (rule subset_inj_on) blast
-    moreover
-    have \<open>inj_on g' (- A)\<close>
-      unfolding g'_def
-    proof (rule inj_on_inv_into)
-      have \<open>g ` (- (f ` A)) \<subseteq> range g\<close> by blast
-      then show \<open>- A \<subseteq> range g\<close> by (simp only: A_compl)
-    qed
-    moreover
-    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
-    proof -
-      from a have fa: \<open>f a \<in> f ` A\<close> by (rule imageI)
-      from b have \<open>g' b \<in> g' ` (- A)\<close> by (rule imageI)
-      with * have \<open>g' b \<in> - (f ` A)\<close> by simp
-      with eq fa show \<open>False\<close> by simp
-    qed
-    ultimately show \<open>inj ?h\<close>
-      unfolding inj_on_def by (metis ComplI)
-  qed
-qed
-
-end
--- a/src/HOL/Library/Extended_Nonnegative_Real.thy	Sat Oct 01 15:21:43 2016 +0200
+++ b/src/HOL/Library/Extended_Nonnegative_Real.thy	Sat Oct 01 20:03:17 2016 +0200
@@ -99,7 +99,7 @@
        (auto intro!: mono_funpow sup_continuous_mono[OF f] SUP_least)
 
   show "lfp g \<le> \<alpha> (lfp f)"
-    by (rule lfp_lowerbound) (simp add: eq[symmetric] lfp_unfold[OF mf, symmetric])
+    by (rule lfp_lowerbound) (simp add: eq[symmetric] lfp_fixpoint[OF mf])
 qed
 
 lemma sup_continuous_applyD: "sup_continuous f \<Longrightarrow> sup_continuous (\<lambda>x. f x h)"
--- a/src/HOL/Library/Order_Continuity.thy	Sat Oct 01 15:21:43 2016 +0200
+++ b/src/HOL/Library/Order_Continuity.thy	Sat Oct 01 20:03:17 2016 +0200
@@ -123,7 +123,7 @@
       case (Suc i)
       have "(F ^^ Suc i) bot = F ((F ^^ i) bot)" by simp
       also have "\<dots> \<le> F (lfp F)" by (rule monoD[OF mono Suc])
-      also have "\<dots> = lfp F" by (simp add: lfp_unfold[OF mono, symmetric])
+      also have "\<dots> = lfp F" by (simp add: lfp_fixpoint[OF mono])
       finally show ?case .
     qed simp
   qed
@@ -287,7 +287,7 @@
     fix i show "gfp F \<le> (F ^^ i) top"
     proof (induct i)
       case (Suc i)
-      have "gfp F = F (gfp F)" by (simp add: gfp_unfold[OF mono, symmetric])
+      have "gfp F = F (gfp F)" by (simp add: gfp_fixpoint[OF mono])
       also have "\<dots> \<le> F ((F ^^ i) top)" by (rule monoD[OF mono Suc])
       also have "\<dots> = (F ^^ Suc i) top" by simp
       finally show ?case .
--- a/src/HOL/Nat.thy	Sat Oct 01 15:21:43 2016 +0200
+++ b/src/HOL/Nat.thy	Sat Oct 01 20:03:17 2016 +0200
@@ -1554,7 +1554,7 @@
       by (simp add: comp_def)
   qed
   have "(f ^^ n) (lfp f) = lfp f" for n
-    by (induct n) (auto intro: f lfp_unfold[symmetric])
+    by (induct n) (auto intro: f lfp_fixpoint)
   then show "lfp (f ^^ Suc n) \<le> lfp f"
     by (intro lfp_lowerbound) (simp del: funpow.simps)
 qed
@@ -1571,7 +1571,7 @@
       by (simp add: comp_def)
   qed
   have "(f ^^ n) (gfp f) = gfp f" for n
-    by (induct n) (auto intro: f gfp_unfold[symmetric])
+    by (induct n) (auto intro: f gfp_fixpoint)
   then show "gfp (f ^^ Suc n) \<ge> gfp f"
     by (intro gfp_upperbound) (simp del: funpow.simps)
 qed
--- a/src/HOL/ROOT	Sat Oct 01 15:21:43 2016 +0200
+++ b/src/HOL/ROOT	Sat Oct 01 20:03:17 2016 +0200
@@ -644,7 +644,6 @@
     Peirce
     Drinker
     Cantor
-    Schroeder_Bernstein
     Structured_Statements
     Basic_Logic
     Expr_Compiler
--- a/src/HOL/Wellfounded.thy	Sat Oct 01 15:21:43 2016 +0200
+++ b/src/HOL/Wellfounded.thy	Sat Oct 01 20:03:17 2016 +0200
@@ -837,8 +837,9 @@
       qed
     qed
   next
-    case [simp]: infinite
-    show ?case by (rule accI) (auto elim: max_ext.cases)
+    case infinite
+    show ?case
+      by (rule accI) (auto elim: max_ext.cases simp: infinite)
   qed
 qed