# HG changeset patch # User wenzelm # Date 1475344997 -7200 # Node ID db9259a5e20eb9af0c302abef6394ae8c8d504b9 # Parent efc958d2fe007d29f1de67a6cfc82d565a5790e3# Parent 4c4049e3bad8739446fc71b8a3f10437b28e6da3 merged diff -r efc958d2fe00 -r db9259a5e20e NEWS --- 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. diff -r efc958d2fe00 -r db9259a5e20e src/HOL/BNF_Cardinal_Order_Relation.thy --- 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| \o |B|" using ordIso_iff_ordLeq by auto } diff -r efc958d2fe00 -r db9259a5e20e src/HOL/Complete_Partial_Order.thy --- 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 \) (op \) f" + from mono have f': "monotone (op \) (op \) f" unfolding mono_def monotone_def . show "lfp f \ fixp f" by (rule lfp_lowerbound, subst fixp_unfold [OF f'], rule order_refl) show "fixp f \ 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 diff -r efc958d2fe00 -r db9259a5e20e src/HOL/Finite_Set.thy --- 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 \ (\(n::nat) f. A = f ` {i::nat. i < n})" +lemma finite_conv_nat_seg_image: "finite A \ (\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 "\f n::nat. f ` A = {i. i < n} \ inj_on f A" + shows "\f n. f ` A = {i::nat. i < n} \ inj_on f A" proof - from finite_imp_nat_seg_image_inj_on [OF \finite A\] obtain f and n :: nat where bij: "bij_betw f {i. iThe Cantor-Bernstein Theorem\ - -lemma Cantor_Bernstein_aux: - "\A' h. A' \ A \ - (\a \ A'. a \ g ` (B - f ` A')) \ - (\a \ A'. h a = f a) \ - (\a \ A - A'. h a \ B - (f ` A') \ 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' \ A" - and 2: "\a \ A'. a \ g ` (B - f ` A')" - and 3: "\a \ A - A'. \b \ B - (f ` A'). a = g b" - by blast+ - define h where "h a = (if a \ A' then f a else (SOME b. b \ B - (f ` A') \ a = g b))" for a - then have 4: "\a \ A'. h a = f a" by simp - have "\a \ A - A'. h a \ B - (f ` A') \ a = g (h a)" - proof - fix a - let ?phi = "\b. b \ B - (f ` A') \ a = g b" - assume *: "a \ A - A'" - from * have "h a = (SOME b. ?phi b)" by (auto simp: h_def) - moreover from 3 * have "\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 \ B" - and inj2: "inj_on g B" and sub2: "g ` B \ A" - shows "\h. bij_betw h A B" -proof- - obtain A' and h where "A' \ A" - and 1: "\a \ A'. a \ g ` (B - f ` A')" - and 2: "\a \ A'. h a = f a" - and 3: "\a \ A - A'. h a \ B - (f ` A') \ 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 \ A" and 5: "a2 \ A" and 6: "h a1 = h a2" - show "a1 = a2" - proof (cases "a1 \ A'") - case True - show ?thesis - proof (cases "a2 \ A'") - case True': True - with True 2 6 have "f a1 = f a2" by auto - with inj1 \A' \ A\ 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 \ 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 \ A" - with sub1 2 3 show "h a \ B" by (cases "a \ A'") auto - next - fix b - assume *: "b \ B" - show "b \ h ` A" - proof (cases "b \ f ` A'") - case True - then obtain a where "a \ A'" "b = f a" by blast - with \A' \ A\ 2 show ?thesis by force - next - case False - with 1 * have "g b \ A'" by auto - with sub2 * have 4: "g b \ A - A'" by auto - with 3 have "h (g b) \ 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 \Other Consequences of Hilbert's Epsilon\ text \Hilbert's Epsilon and the @{term split} Operator\ @@ -566,30 +466,48 @@ text \A relation is wellfounded iff it has no infinite descending chain.\ -lemma wf_iff_no_infinite_down_chain: "wf r \ (\ (\f. \i. (f (Suc i), f i) \ r))" - apply (simp only: wf_eq_minimal) - apply (rule iffI) - apply (rule notI) - apply (erule exE) - apply (erule_tac x = "{w. \i. w = f i}" in allE) - apply blast - apply (erule contrapos_np) - apply simp - apply clarify - apply (subgoal_tac "\n. rec_nat x (\i y. SOME z. z \ Q \ (z, y) \ r) n \ Q") - apply (rule_tac x = "rec_nat x (\i y. SOME z. z \ Q \ (z, y) \ 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 \ (\f. \i. (f (Suc i), f i) \ r)" + (is "_ \ \ ?ex") +proof + assume "wf r" + show "\ ?ex" + proof + assume ?ex + then obtain f where f: "(f (Suc i), f i) \ r" for i + by blast + from \wf r\ have minimal: "x \ Q \ \z\Q. \y. (y, z) \ r \ y \ Q" for x Q + by (auto simp: wf_eq_minimal) + let ?Q = "{w. \i. w = f i}" + fix n + have "f n \ ?Q" by blast + from minimal [OF this] obtain j where "(y, f j) \ r \ y \ ?Q" for y by blast + with this [OF \(f (Suc j), f j) \ r\] have "f (Suc j) \ ?Q" by simp + then show False by blast + qed +next + assume "\ ?ex" + then show "wf r" + proof (rule contrapos_np) + assume "\ wf r" + then obtain Q x where x: "x \ Q" and rec: "z \ Q \ \y. (y, z) \ r \ y \ Q" for z + by (auto simp add: wf_eq_minimal) + obtain descend :: "nat \ 'a" + where descend_0: "descend 0 = x" + and descend_Suc: "descend (Suc n) = (SOME y. y \ Q \ (y, descend n) \ r)" for n + by (rule that [of "rec_nat x (\_ rec. (SOME y. y \ Q \ (y, rec) \ r))"]) simp_all + have descend_Q: "descend n \ 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) \ r" for i + by (simp only: descend_Suc) (rule someI2_ex; use descend_Q rec in blast) + then show "\f. \i. (f (Suc i), f i) \ r" by blast + qed +qed lemma wf_no_infinite_down_chainE: assumes "wf r" diff -r efc958d2fe00 -r db9259a5e20e src/HOL/Inductive.thy --- 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 \Least and greatest fixed points\ +subsection \Least fixed points\ context complete_lattice begin -definition lfp :: "('a \ 'a) \ 'a" \ \least fixed point\ +definition lfp :: "('a \ 'a) \ 'a" where "lfp f = Inf {u. f u \ u}" -definition gfp :: "('a \ 'a) \ 'a" \ \greatest fixed point\ - where "gfp f = Sup {u. u \ f u}" - - -subsection \Proof of Knaster-Tarski Theorem using @{term lfp}\ - -text \@{term "lfp f"} is the least upper bound of the set @{term "{u. f u \ u}"}\ - lemma lfp_lowerbound: "f A \ A \ lfp f \ A" - by (auto simp add: lfp_def intro: Inf_lower) + unfolding lfp_def by (rule Inf_lower) simp lemma lfp_greatest: "(\u. f u \ u \ A \ u) \ A \ 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 \ f (lfp f) \ lfp f" - by (iprover intro: lfp_greatest order_trans monoD lfp_lowerbound) - -lemma lfp_lemma3: "mono f \ lfp f \ 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 \ u}" + let ?a = "\?H" + show "f ?a \ ?a" + proof (rule Inf_greatest) + fix x + assume "x \ ?H" + then have "?a \ x" by (rule Inf_lower) + with \mono f\ have "f ?a \ f x" .. + also from \x \ ?H\ have "f x \ x" .. + finally show "f ?a \ x" . + qed + show "?a \ f ?a" + proof (rule Inf_lower) + from \mono f\ and \f ?a \ ?a\ have "f (f ?a) \ f ?a" .. + then show "f ?a \ ?H" .. + qed +qed lemma lfp_unfold: "mono f \ lfp f = f (lfp f)" - by (iprover intro: order_antisym lfp_lemma2 lfp_lemma3) + by (rule lfp_fixpoint [symmetric]) lemma lfp_const: "lfp (\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 \Proof of Knaster-Tarski Theorem using \gfp\\ +subsection \Greatest fixed points\ -text \@{term "gfp f"} is the greatest lower bound of the set @{term "{u. u \ f u}"}\ +context complete_lattice +begin + +definition gfp :: "('a \ 'a) \ 'a" + where "gfp f = Sup {u. u \ f u}" lemma gfp_upperbound: "X \ f X \ X \ gfp f" by (auto simp add: gfp_def intro: Sup_upper) @@ -142,14 +155,36 @@ lemma gfp_least: "(\u. u \ f u \ u \ X) \ gfp f \ X" by (auto simp add: gfp_def intro: Sup_least) -lemma gfp_lemma2: "mono f \ gfp f \ f (gfp f)" - by (iprover intro: gfp_least order_trans monoD gfp_upperbound) +end + +lemma lfp_le_gfp: "mono f \ lfp f \ gfp f" + by (rule gfp_upperbound) (simp add: lfp_fixpoint) -lemma gfp_lemma3: "mono f \ f (gfp f) \ 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 \ f u}" + let ?a = "\?H" + show "?a \ f ?a" + proof (rule Sup_least) + fix x + assume "x \ ?H" + then have "x \ f x" .. + also from \x \ ?H\ have "x \ ?a" by (rule Sup_upper) + with \mono f\ have "f x \ f ?a" .. + finally show "x \ f ?a" . + qed + show "f ?a \ ?a" + proof (rule Sup_upper) + from \mono f\ and \?a \ f ?a\ have "f ?a \ f (f ?a)" .. + then show "f ?a \ ?H" .. + qed +qed lemma gfp_unfold: "mono f \ gfp f = f (gfp f)" - by (iprover intro: order_antisym gfp_lemma2 gfp_lemma3) + by (rule gfp_fixpoint [symmetric]) lemma gfp_const: "gfp (\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 \ lfp f \ gfp f" - by (iprover intro: gfp_upperbound lfp_lemma3) - - subsection \Coinduction rules for greatest fixed points\ text \Weak version.\ @@ -174,7 +205,7 @@ done lemma coinduct_lemma: "X \ f (sup X (gfp f)) \ mono f \ sup X (gfp f) \ 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 \ a \ gfp f \ a \ f (X \ 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 \ 'a" @@ -248,7 +279,7 @@ "X \ f (lfp (\x. f x \ X \ gfp f)) \ mono f \ lfp (\x. f x \ X \ gfp f) \ f (lfp (\x. f x \ X \ 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 \The Schroeder-Bernstein Theorem\ + +text \ + See also: + \<^item> \<^file>\$ISABELLE_HOME/src/HOL/ex/Set_Theory.thy\ + \<^item> \<^url>\http://planetmath.org/proofofschroederbernsteintheoremusingtarskiknastertheorem\ + \<^item> Springer LNCS 828 (cover page) +\ + +theorem Schroeder_Bernstein: + fixes f :: "'a \ 'b" and g :: "'b \ 'a" + and A :: "'a set" and B :: "'b set" + assumes inj1: "inj_on f A" and sub1: "f ` A \ B" + and inj2: "inj_on g B" and sub2: "g ` B \ A" + shows "\h. bij_betw h A B" +proof (rule exI, rule bij_betw_imageI) + define X where "X = lfp (\X. A - (g ` (B - (f ` X))))" + define g' where "g' = the_inv_into (B - (f ` X)) g" + let ?h = "\z. if z \ 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 \ A" by auto + from X sub1 have fX_sub: "f ` X \ B" by auto + + show "?h ` A = B" + proof - + from X_sub have "?h ` A = ?h ` (X \ (A - X))" by auto + also have "\ = ?h ` X \ ?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 \ (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 \ X" and b: "b \ A - X" for a b + proof - + from a have fa: "f a \ f ` X" by (rule imageI) + from b have "g' b \ g' ` (A - X)" by (rule imageI) + with * have "g' b \ - (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 \ A" and "b \ A" + then consider "a \ X" "b \ X" | "a \ A - X" "b \ A - X" + | "a \ X" "b \ A - X" | "a \ A - X" "b \ 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 \Inductive datatypes and primitive recursion\ text \Package setup.\ diff -r efc958d2fe00 -r db9259a5e20e src/HOL/Isar_Examples/Schroeder_Bernstein.thy --- 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 \Schröder-Bernstein Theorem\ - -theory Schroeder_Bernstein - imports Main -begin - -text \ - See also: - \<^item> \<^file>\$ISABELLE_HOME/src/HOL/ex/Set_Theory.thy\ - \<^item> \<^url>\http://planetmath.org/proofofschroederbernsteintheoremusingtarskiknastertheorem\ - \<^item> Springer LNCS 828 (cover page) -\ - -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\ - - have \A = - (g ` (- (f ` A)))\ - unfolding A_def by (rule lfp_unfold) (blast intro: monoI) - then have A_compl: \- A = g ` (- (f ` A))\ by blast - then have *: \g' ` (- A) = - (f ` A)\ - using g'_def \inj g\ by auto - - show \inj ?h \ surj ?h\ - proof - from * show \surj ?h\ by auto - have \inj_on f A\ - using \inj f\ by (rule subset_inj_on) blast - moreover - have \inj_on g' (- A)\ - unfolding g'_def - proof (rule inj_on_inv_into) - have \g ` (- (f ` A)) \ range g\ by blast - 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 - 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 - qed - ultimately show \inj ?h\ - unfolding inj_on_def by (metis ComplI) - qed -qed - -end diff -r efc958d2fe00 -r db9259a5e20e src/HOL/Library/Extended_Nonnegative_Real.thy --- 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 \ \ (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 \ sup_continuous (\x. f x h)" diff -r efc958d2fe00 -r db9259a5e20e src/HOL/Library/Order_Continuity.thy --- 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 "\ \ F (lfp F)" by (rule monoD[OF mono Suc]) - also have "\ = lfp F" by (simp add: lfp_unfold[OF mono, symmetric]) + also have "\ = lfp F" by (simp add: lfp_fixpoint[OF mono]) finally show ?case . qed simp qed @@ -287,7 +287,7 @@ fix i show "gfp F \ (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 "\ \ F ((F ^^ i) top)" by (rule monoD[OF mono Suc]) also have "\ = (F ^^ Suc i) top" by simp finally show ?case . diff -r efc958d2fe00 -r db9259a5e20e src/HOL/Nat.thy --- 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) \ 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) \ gfp f" by (intro gfp_upperbound) (simp del: funpow.simps) qed diff -r efc958d2fe00 -r db9259a5e20e src/HOL/ROOT --- 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 diff -r efc958d2fe00 -r db9259a5e20e src/HOL/Wellfounded.thy --- 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