# HG changeset patch # User paulson # Date 1744927046 -3600 # Node ID 62afd98e3f3e25e59ad0a7274dc19d9ac733faed # Parent 819688d4cc45a741107aa920aa496784bec210ae more tidying diff -r 819688d4cc45 -r 62afd98e3f3e src/HOL/Analysis/Elementary_Topology.thy --- a/src/HOL/Analysis/Elementary_Topology.thy Wed Apr 16 21:13:33 2025 +0100 +++ b/src/HOL/Analysis/Elementary_Topology.thy Thu Apr 17 22:57:26 2025 +0100 @@ -54,13 +54,8 @@ lemma topological_basis: "topological_basis B \ (\x. open x \ (\B'. B' \ B \ \B' = x))" (is "?lhs = ?rhs") -proof - show "?lhs \ ?rhs" - by (meson local.open_Union subsetD topological_basis_def) - show "?rhs \ ?lhs" - unfolding topological_basis_def - by (metis cSup_singleton empty_subsetI insert_subset) -qed + by (metis bot.extremum cSup_singleton insert_subset open_Union subsetD + topological_basis_def) lemma topological_basis_iff: assumes "\B'. B' \ B \ open B'" @@ -69,18 +64,19 @@ proof safe fix O' and x::'a assume H: "topological_basis B" "open O'" "x \ O'" - then have "(\B'\B. \B' = O')" by (simp add: topological_basis_def) - then obtain B' where "B' \ B" "O' = \B'" by auto - then show "\B'\B. x \ B' \ B' \ O'" using H by auto + then obtain B' where "B'\B" "\B' = O'" + by (force simp add: topological_basis_def) + then show "\B'\B. x \ B' \ B' \ O'" + using H by auto next assume H: ?rhs show "topological_basis B" - using assms unfolding topological_basis_def - proof safe + unfolding topological_basis_def + proof (intro conjI strip assms) fix O' :: "'a set" assume "open O'" with H obtain f where "\x\O'. f x \ B \ x \ f x \ f x \ O'" - by (force intro: bchoice simp: Bex_def) + by metis then show "\B'\B. \B' = O'" by (auto intro: exI[where x="{f x |x. x \ O'}"]) qed @@ -156,9 +152,9 @@ then show ?thesis using subset_eq by force qed - with A B show ?thesis + with A B open_Times show ?thesis unfolding topological_basis_def - by (smt (verit) SigmaE imageE image_mono open_Times case_prod_conv) + by (smt (verit) SigmaE imageE image_mono case_prod_conv) qed @@ -231,7 +227,7 @@ "\ = (\N. \((\n. from_nat_into \ n) ` N)) ` (Collect finite::nat set set)" then show "\\. countable \ \ (\A. A \ \ \ x \ A) \ (\A. A \ \ \ open A) \ (\S. open S \ x \ S \ (\A\\. A \ S)) \ (\A B. A \ \ \ B \ \ \ A \ B \ \)" - proof (safe intro!: exI[where x=\]) + proof (intro exI conjI strip) show "countable \" unfolding \_def by (intro countable_image countable_Collect_finite) fix A @@ -266,14 +262,14 @@ and 1: "\A. A \ \ \ x \ A" "\A. A \ \ \ open A" and 2: "\S. open S \ x \ S \ \A\\. A \ S" shows "\\::nat \ 'a set. (\i. x \ \ i \ open (\ i)) \ (\S. open S \ x \ S \ (\i. \ i \ S))" -proof (safe intro!: exI[of _ "from_nat_into \"]) +proof (intro exI[of _ "from_nat_into _"] conjI strip) fix i have "\ \ {}" using 2[of UNIV] by auto show "x \ from_nat_into \ i" "open (from_nat_into \ i)" using range_from_nat_into_subset[OF \\ \ {}\] 1 by auto next fix S - assume "open S" "x\S" + assume "open S \ x\S" then show "\i. from_nat_into \ i \ S" by (metis "2" \countable \\ from_nat_into_surj) qed @@ -331,11 +327,10 @@ proof (intro exI conjI) show "countable ?B" by (intro countable_image countable_Collect_finite_subset B) - { - fix S - assume "open S" - then have "\B'\{b. finite b \ b \ B}. (\b\B'. \b) = S" - unfolding B + have "\B'\Inter ` {b. finite b \ b \ B}. \B' = S" if "open S" for S + proof - + have "\B'\{b. finite b \ b \ B}. (\b\B'. \b) = S" + using that unfolding B proof induct case UNIV show ?case by (intro exI[of _ "{{}}"]) simp @@ -360,12 +355,12 @@ then show ?case by (intro exI[of _ "{{S}}"]) auto qed - then have "(\B'\Inter ` {b. finite b \ b \ B}. \B' = S)" - unfolding subset_image_iff by blast } + then show ?thesis + unfolding subset_image_iff by blast + qed then show "topological_basis ?B" unfolding topological_basis_def - by (safe intro!: open_Inter) - (simp_all add: B generate_topology.Basis subset_eq) + by (safe intro!: open_Inter) (simp_all add: B generate_topology.Basis subset_eq) qed qed @@ -394,14 +389,11 @@ by (simp add: \_def) then obtain G where G: "\S. S \ \ \ G S \ \ \ S \ G S" by metis - have "\\ \ \\" + have eq1: "\\ = \\" unfolding \_def by (blast dest: \ \) - moreover have "\\ \ \\" - using \_def by blast - ultimately have eq1: "\\ = \\" .. - moreover have "\\ = \ (G ` \)" + also have "\ = \ (G ` \)" using G eq1 by auto - ultimately show ?thesis + finally show ?thesis by (metis G \countable \\ countable_image image_subset_iff that) qed @@ -464,33 +456,30 @@ define B2 where "B2 = {(SOME x. x \ U)| U. U \ A}" then have "countable B2" using \countable A\ by (simp add: Setcompr_eq_image) have "\b \ B1 \ B2. x < b \ b \ y" if "x < y" for x y - proof (cases) - assume "\z. x < z \ z < y" + proof (cases "\z. x < z \ z < y") + case True then obtain z where z: "x < z \ z < y" by auto define U where "U = {x<.. U" using z U_def by simp - ultimately obtain V where "V \ A" "z \ V" "V \ U" - using topological_basisE[OF \topological_basis A\] by auto + then obtain V where "V \ A" "z \ V" "V \ U" + using topological_basisE[OF \topological_basis A\] + by (metis U_def greaterThanLessThan_iff z) define w where "w = (SOME x. x \ V)" then have "w \ V" using \z \ V\ by (metis someI2) - then have "x < w \ w \ y" using \w \ V\ \V \ U\ U_def by fastforce - moreover have "w \ B1 \ B2" using w_def B2_def \V \ A\ by auto - ultimately show ?thesis by auto + with \V \ A\ \V \ U\ show ?thesis + by (force simp: B2_def U_def w_def) next - assume "\(\z. x < z \ z < y)" + case False then have *: "\z. z > x \ z \ y" by auto define U where "U = {x<..}" then have "open U" by simp - moreover have "y \ U" using \x < y\ U_def by simp - ultimately obtain "V" where "V \ A" "y \ V" "V \ U" - using topological_basisE[OF \topological_basis A\] by auto + then obtain "V" where "V \ A" "y \ V" "V \ U" + using topological_basisE[OF \topological_basis A\] U_def \x < y\ by blast have "U = {y..}" unfolding U_def using * \x < y\ by auto then have "V \ {y..}" using \V \ U\ by simp then have "(LEAST w. w \ V) = y" using \y \ V\ by (meson Least_equality atLeast_iff subsetCE) - then have "y \ B1 \ B2" using \V \ A\ B1_def by auto - moreover have "x < y \ y \ y" using \x < y\ by simp - ultimately show ?thesis by auto + then show ?thesis + using B1_def \V \ A\ that by blast qed moreover have "countable (B1 \ B2)" using \countable B1\ \countable B2\ by simp ultimately show ?thesis by auto @@ -505,33 +494,35 @@ define B2 where "B2 = {(SOME x. x \ U)| U. U \ A}" then have "countable B2" using \countable A\ by (simp add: Setcompr_eq_image) have "\b \ B1 \ B2. x \ b \ b < y" if "x < y" for x y - proof (cases) - assume "\z. x < z \ z < y" + proof (cases "\z. x < z \ z < y") + case True then obtain z where z: "x < z \ z < y" by auto define U where "U = {x<.. U" using z U_def by simp - ultimately obtain "V" where "V \ A" "z \ V" "V \ U" - using topological_basisE[OF \topological_basis A\] by auto + then obtain "V" where "V \ A" "z \ V" "V \ U" + using topological_basisE[OF \topological_basis A\] + by (metis U_def greaterThanLessThan_iff z) define w where "w = (SOME x. x \ V)" - then have "w \ V" using \z \ V\ by (metis someI2) - then have "x \ w \ w < y" using \w \ V\ \V \ U\ U_def by fastforce - moreover have "w \ B1 \ B2" using w_def B2_def \V \ A\ by auto - ultimately show ?thesis by auto + then have "w \ V" + using \z \ V\ by (metis someI2) + then have "x \ w \ w < y" + using \w \ V\ \V \ U\ U_def by fastforce + then show ?thesis + using B2_def \V \ A\ w_def by blast next - assume "\(\z. x < z \ z < y)" + case False then have *: "\z. z < y \ z \ x" using leI by blast define U where "U = {.. U" using \x < y\ U_def by simp - ultimately obtain "V" where "V \ A" "x \ V" "V \ U" - using topological_basisE[OF \topological_basis A\] by auto + then obtain "V" where "V \ A" "x \ V" "V \ U" + using topological_basisE[OF \topological_basis A\] U_def \x < y\ by blast have "U = {..x}" unfolding U_def using * \x < y\ by auto - then have "V \ {..x}" using \V \ U\ by simp - then have "(GREATEST x. x \ V) = x" using \x \ V\ by (meson Greatest_equality atMost_iff subsetCE) - then have "x \ B1 \ B2" using \V \ A\ B1_def by auto - moreover have "x \ x \ x < y" using \x < y\ by simp - ultimately show ?thesis by auto + then have "V \ {..x}" + using \V \ U\ by simp + then have "(GREATEST x. x \ V) = x" + using \x \ V\ by (meson Greatest_equality atMost_iff subsetCE) + then show ?thesis + using B1_def \V \ A\ that by blast qed moreover have "countable (B1 \ B2)" using \countable B1\ \countable B2\ by simp ultimately show ?thesis by auto @@ -546,8 +537,8 @@ proof - obtain z where "x < z" "z < y" using \x < y\ dense by blast then obtain b where "b \ B" "x < b \ b \ z" using B(2) by auto - then have "x < b \ b < y" using \z < y\ by auto - then show ?thesis using \b \ B\ by auto + then show ?thesis + using \z < y\ by auto qed then show ?thesis using B(1) by auto qed @@ -686,14 +677,11 @@ with A(3) have "eventually (\i. A i \ S) sequentially" by auto moreover - { - fix i - assume "Suc 0 \ i" - then have "f (r i) \ A i" + have "eventually (\i. f (r i) \ A i) sequentially" + proof + fix i assume "Suc 0 \ i" then show "f (r i) \ A i" by (cases i) (simp_all add: r_def s) - } - then have "eventually (\i. f (r i) \ A i) sequentially" - by (auto simp: eventually_sequentially) + qed ultimately show "eventually (\i. f (r i) \ S) sequentially" by eventually_elim auto qed @@ -710,8 +698,7 @@ lemma sequence_unique_limpt: fixes f :: "nat \ 'a::t2_space" - assumes f: "(f \ l) sequentially" - and "l' islimpt (range f)" + assumes f: "(f \ l) sequentially" and l': "l' islimpt (range f)" shows "l' = l" proof (rule ccontr) assume "l' \ l" @@ -722,16 +709,12 @@ have "UNIV = {.. {N..}" by auto - then have "l' islimpt (f ` ({.. {N..}))" - using assms(2) by simp then have "l' islimpt (f ` {.. f ` {N..})" - by (simp add: image_Un) + by (metis l' image_Un) then have "l' islimpt (f ` {N..})" by (simp add: islimpt_Un_finite) then obtain y where "y \ f ` {N..}" "y \ s" "y \ l'" using \l' \ s\ \open s\ by (rule islimptE) - then obtain n where "N \ n" "f n \ s" "f n \ l'" - by auto with \\n\N. f n \ t\ \s \ t = {}\ show False by blast qed @@ -806,7 +789,8 @@ lemma islimpt_image: assumes "z islimpt g -` A \ B" "g z \ A" "z \ B" "continuous_on B g" shows "g z islimpt A" - by (smt (verit, best) IntD1 assms continuous_on_topological inf_le2 islimpt_def subset_eq vimageE) + using assms + by (simp add: islimpt_def vimage_def continuous_on_topological Bex_def) metis subsection \Interior of a Set\ @@ -865,7 +849,10 @@ by (meson interior_eq interior_subset not_open_singleton subset_singletonD) lemma interior_Int [simp]: "interior (S \ T) = interior S \ interior T" - by (meson Int_mono Int_subset_iff antisym_conv interior_maximal interior_subset open_Int open_interior) +proof + show "interior S \ interior T \ interior (S \ T)" + by (meson Int_mono interior_subset open_Int open_interior open_subset_interior) +qed (simp add: interior_mono) lemma eventually_nhds_in_nhd: "x \ interior s \ eventually (\y. y \ s) (nhds x)" using interior_subset[of s] by (subst eventually_nhds) blast @@ -913,14 +900,14 @@ proof fix x assume "x \ interior (S \ T)" - then obtain R where "open R" "x \ R" "R \ S \ T" .. + then obtain R where R: "open R" "x \ R" "R \ S \ T" .. show "x \ interior S" proof (rule ccontr) assume "x \ interior S" with \x \ R\ \open R\ obtain y where "y \ R - S" unfolding interior_def by fast - then show False - by (metis Diff_subset_conv \R \ S \ T\ \open R\ cS empty_iff iT interiorI open_Diff) + with R show False + by (metis Diff_subset_conv cS empty_iff iT interiorI open_Diff) qed qed qed @@ -996,9 +983,9 @@ then show "countable (interior ` \)" by (auto intro: countable_disjoint_open_subsets) show "inj_on interior \" - using pw apply (clarsimp simp: inj_on_def pairwise_def) - apply (metis disjnt_def disjnt_subset1 inf.orderE int interior_subset) - done + using pw + unfolding inj_on_def pairwise_def disjnt_def + by (metis inf.idem int interior_Int interior_empty) qed subsection \Closure of a Set\ @@ -1089,21 +1076,28 @@ show "closed (closure A \ closure B)" by (intro closed_Times closed_closure) fix T - assume "A \ B \ T" and "closed T" + assume T: "A \ B \ T" "closed T" + have False + if ab: "a \ closure A" "b \ closure B" and "(a, b) \ T" for a b + proof - + obtain C D where "open C" "open D" "C \ D \ - T" "a \ C" "b \ D" + by (metis ComplI SigmaE2 \closed T\ \(a, b) \ T\ open_Compl open_prod_elim) + then obtain "\ C \ - A" "\ D \ - B" + by (meson ab disjoint_iff inf_shunt open_Int_closure_eq_empty) + then show False + using T \C \ D \ - T\ by auto + qed then show "closure A \ closure B \ T" - apply (simp add: closed_def open_prod_def, clarify) - apply (rule ccontr) - apply (drule_tac x="(a, b)" in bspec, simp, clarify, rename_tac C D) - apply (simp add: closure_interior interior_def) - apply (drule_tac x=C in spec) - apply (drule_tac x=D in spec, auto) - done + by blast qed lemma closure_open_Int_superset: assumes "open S" "S \ closure T" shows "closure(S \ T) = closure S" - by (metis Int_Un_distrib Un_Int_eq(4) assms closure_Un closure_closure open_Int_closure_subset sup.orderE) +proof + show "closure S \ closure (S \ T)" + by (metis assms closed_closure closure_minimal inf.absorb_iff1 open_Int_closure_subset) +qed (simp add: closure_mono) lemma closure_Int: "closure (\I) \ \{closure S |S. S \ I}" by (simp add: INF_greatest Inter_lower Setcompr_eq_image closure_mono) @@ -1154,7 +1148,7 @@ lemma frontier_Int_closed: assumes "closed S" "closed T" shows "frontier(S \ T) = (frontier S \ T) \ (S \ frontier T)" - by (smt (verit, best) Diff_Int Int_Diff assms closed_Int closure_eq frontier_def inf_commute interior_Int) + by (simp add: Int_Un_distrib assms closed_Int frontier_closures inf_commute inf_left_commute) lemma frontier_subset_closed: "closed S \ frontier S \ S" by (metis frontier_def closure_closed Diff_subset) @@ -1209,7 +1203,7 @@ lemma not_in_closure_trivial_limitI: "x \ closure S \ trivial_limit (at x within S)" using not_trivial_limit_within[of x S] - by safe (metis Diff_empty Diff_insert0 closure_subset contra_subsetD) + by (metis Diff_empty Diff_insert0 closure_subset subsetD) lemma filterlim_at_within_closure_implies_filterlim: "filterlim f l (at x within S)" if "x \ closure S \ filterlim f l (at x within S)" @@ -1322,7 +1316,7 @@ lemma closure_sequential: fixes l :: "'a::first_countable_topology" shows "l \ closure S \ (\x. (\n. x n \ S) \ (x \ l) sequentially)" - by (metis Diff_empty Diff_insert0 Un_iff closure_def islimpt_sequential mem_Collect_eq tendsto_const) + by (auto simp: closure_def islimpt_sequential) lemma closed_sequential_limits: fixes S :: "'a::first_countable_topology set" @@ -1340,13 +1334,12 @@ have *: "(S \ T) \ {x. x \ S} = S" "(S \ T) \ {x. x \ S} = T - S" by auto have "(f \ l x) (at x within S)" - by (rule filterlim_at_within_closure_implies_filterlim) - (use \x \ _\ in \auto simp: inf_commute closure_def intro: tendsto_within_subset[OF f]\) + using tendsto_within_subset[OF f] \x \ S \ T\ + by (metis Int_iff Un_iff Un_upper1 closure_def filterlim_at_within_closure_implies_filterlim) moreover have "(g \ l x) (at x within T - S)" - by (rule filterlim_at_within_closure_implies_filterlim) - (use \x \ _\ in - \auto intro!: tendsto_within_subset[OF g] simp: closure_def intro: islimpt_subset\) + using tendsto_within_subset g \x \ S \ T\ + by (metis IntI Un_Diff_Int Un_iff Un_upper1 closure_def filterlim_at_within_closure_implies_filterlim) ultimately show ?thesis by (intro filterlim_at_within_If) (simp_all only: *) qed @@ -1393,12 +1386,12 @@ then have g': "\x\g. \y \ S. x = f y" by auto have "inj_on f T" - by (smt (verit, best) assms(3) f inj_onI subsetD) + unfolding inj_on_def using \T \ S\ f by blast then have "infinite (f ` T)" using assms(2) using finite_imageD by auto moreover - have False if "x \ T" "f x \ g" for x - by (smt (verit) UnionE assms(3) f g' g(3) subsetD that) + have False if "x \ T" "f x \ g" for x + using \T \ S\ f g' \S \ \g\ that by force then have "f ` T \ g" by auto ultimately show False using g(2) using finite_subset by auto @@ -1406,7 +1399,7 @@ lemma sequence_infinite_lemma: fixes f :: "nat \ 'a::t1_space" - assumes "\n. f n \ l" + assumes "\n. f n \ l" and "(f \ l) sequentially" shows "infinite (range f)" proof @@ -1426,22 +1419,18 @@ proof - { fix x l - assume as: "\n::nat. x n \ S" "(x \ l) sequentially" + assume \
: "\n. x n \ S" "(x \ l) sequentially" have "l \ S" proof (cases "\n. x n \ l") - case False - then show "l\S" using as(1) by auto - next case True - with as(2) have "infinite (range x)" + with \
have "infinite (range x)" using sequence_infinite_lemma[of x l] by auto - then obtain l' where "l'\S" "l' islimpt (range x)" - using as(1) assms by auto - then show "l\S" using sequence_unique_limpt as True by auto - qed + with \
assms show "l\S" + using sequence_unique_limpt \
True by blast + qed (use \
in auto) } then show ?thesis - unfolding closed_sequential_limits by fast + unfolding closed_sequential_limits by auto qed lemma closure_insert: @@ -1513,20 +1502,12 @@ lemma closure_iff_nhds_not_empty: "x \ closure X \ (\A. \S\A. open S \ x \ S \ X \ A \ {})" -proof safe - assume x: "x \ closure X" - fix S A - assume \
: "open S" "x \ S" "X \ A = {}" "S \ A" - then have "x \ closure (-S)" - by (simp add: closed_open) - with x have "x \ closure X - closure (-S)" - by auto - with \
show False - by (metis Compl_iff Diff_iff closure_mono disjoint_iff subsetD subsetI) -next - assume "\A S. S \ A \ open S \ x \ S \ X \ A \ {}" - then show "x \ closure X" - by (metis ComplI Compl_disjoint closure_interior interior_subset open_interior) +proof - + have "\A S. S \ A \ open S \ x \ S \ X \ A \ {} \ x \ closure X" + by (metis ComplI Diff_disjoint Diff_eq closure_interior inf_top_left + interior_subset open_interior) + then show ?thesis + using open_Int_closure_eq_empty by fastforce qed lemma compact_filter: @@ -1595,7 +1576,7 @@ { fix V assume "V \ A" then have "F \ principal V" - unfolding F_def by (intro INF_lower2[of V]) auto + by (metis INF_lower F_def insertCI) then have V: "eventually (\x. x \ V) F" by (auto simp: le_filter_def eventually_principal) have "x \ closure V" @@ -1628,7 +1609,7 @@ using assms unfolding countably_compact_def by metis lemma countably_compactI: - assumes "\C. \t\C. open t \ s \ \C \ countable C \ (\C'\C. finite C' \ s \ \C')" + assumes "\C. \t\C. open t \ s \ \C \ countable C \ \C'\C. finite C' \ s \ \C'" shows "countably_compact s" using assms unfolding countably_compact_def by metis @@ -1651,12 +1632,10 @@ unfolding C_def using ccover by auto moreover have "\A \ U \ \C" - proof safe + proof clarify fix x a assume "x \ U" "x \ a" "a \ A" - with basis[of a x] A obtain b where "b \ B" "x \ b" "b \ U \ a" - by blast - with \a \ A\ show "x \ \C" + with basis[of a x] A show "x \ \C" unfolding C_def by auto qed then have "U \ \C" using \U \ \A\ by auto @@ -1675,10 +1654,9 @@ proof (rule countably_compact_imp_compact) fix T and x :: 'a assume "open T" "x \ T" - from topological_basisE[OF is_basis this] obtain b where - "b \ (SOME B. countable B \ topological_basis B)" "x \ b" "b \ T" . - then show "\b\SOME B. countable B \ topological_basis B. x \ b \ b \ U \ T" - by blast + from topological_basisE[OF is_basis this] + show "\b\SOME B. countable B \ topological_basis B. x \ b \ b \ U \ T" + by (metis le_infI1) qed (insert countable_basis topological_basis_open[OF is_basis], auto) lemma countably_compact_eq_compact: @@ -1726,7 +1704,7 @@ fixes U :: "'a :: first_countable_topology set" assumes "seq_compact U" shows "countably_compact U" -proof (safe intro!: countably_compactI) +proof (intro countably_compactI) fix A assume A: "\a\A. open a" "U \ \A" "countable A" have subseq: "\X. range X \ U \ \r x. x \ U \ strict_mono (r :: nat \ nat) \ (X \ r) \ x" @@ -1824,12 +1802,8 @@ with A(3) have "eventually (\i. A i \ S) sequentially" by auto moreover - { - fix i - assume "Suc 0 \ i" - then have "X (r i) \ A i" - by (cases i) (simp_all add: r_def s) - } + have "X (r i) \ A i" if "Suc 0 \ i" for i + using that by (cases i) (simp_all add: r_def s) then have "eventually (\i. X (r i) \ A i) sequentially" by (auto simp: eventually_sequentially) ultimately show "eventually (\i. X (r i) \ S) sequentially" @@ -1853,13 +1827,11 @@ moreover assume "\ (\x\S. \U. x\U \ open U \ infinite (U \ T))" then have S: "\x. x \ S \ \U. x\U \ open U \ finite (U \ T)" by metis - have "S \ \C" - using \T \ S\ + have "\x U. \T \ S; x \ U; open U; finite (U \ T)\ \ x \ \ C" unfolding C_def - apply (safe dest!: S) - apply (rule_tac a="U \ T" in UN_I) - apply (auto intro!: interiorI simp add: finite_subset) - done + by (auto intro!: UN_I [where a="U \ T"] interiorI simp add: finite_subset) + then have "S \ \C" + using \T \ S\ S by force moreover from \countable T\ have "countable C" unfolding C_def by (auto intro: countable_Collect_finite_subset) @@ -1975,7 +1947,7 @@ moreover from D c have "(\y\D. a y) \ T \ (\y\D. c y)" by (fastforce simp: subset_eq) ultimately show "\a. open a \ x \ a \ (\d\\. finite d \ a \ T \ \d)" - using c by (intro exI[of _ "c`D"] exI[of _ "\(a`D)"] conjI) (auto intro!: open_INT) + using c exI[of _ "c`D"] exI[of _ "\(a`D)"] by (simp add: open_INT subset_eq) qed then obtain a d where a: "\x. x\S \ open (a x)" "S \ (\x\S. a x)" and d: "\x. x \ S \ d x \ \ \ finite (d x) \ a x \ T \ \(d x)" @@ -1988,7 +1960,7 @@ have "S \ T \ (\x\e. \(d x))" by (smt (verit, del_insts) S SigmaE UN_iff d e(1) mem_Sigma_iff subset_eq) ultimately show "\C'\\. finite C' \ S \ T \ \C'" - by (intro exI[of _ "(\x\e. d x)"]) (auto simp: subset_eq) + by (force simp: subset_eq intro!: exI[of _ "\x\e. d x"]) qed @@ -1998,13 +1970,8 @@ assumes "{x0} \ K \ W" shows "\X0. x0 \ X0 \ open X0 \ X0 \ K \ W" proof - - { - fix y assume "y \ K" - then have "(x0, y) \ W" using assms by auto - with \open W\ - have "\X0 Y. open X0 \ open Y \ x0 \ X0 \ y \ Y \ X0 \ Y \ W" - by (rule open_prod_elim) blast - } + have "\X0 Y. open X0 \ open Y \ x0 \ X0 \ y \ Y \ X0 \ Y \ W" if "y \ K" for y + using assms open_prod_def subsetD that by fastforce then obtain X0 Y where *: "\y \ K. open (X0 y) \ open (Y y) \ x0 \ X0 y \ y \ Y y \ X0 y \ Y y \ W" by metis @@ -2014,7 +1981,7 @@ then obtain c where c: "\C. C \ CC \ c C \ K \ C = Y (c C)" by (force intro!: choice) with * CC show ?thesis - by (force intro!: exI[where x="\C\CC. X0 (c C)"]) (* SLOW *) + by (bestsimp intro!: exI[where x="\C\CC. X0 (c C)"]) (* SLOW *) qed lemma continuous_on_prod_compactE: @@ -2047,7 +2014,7 @@ by blast have "\x\X0 \ U. \t \ C. dist (fx (x, t)) (fx (x0, t)) \ e" - proof safe + proof clarify fix x assume x: "x \ X0" "x \ U" fix t assume t: "t \ C" have "dist (fx (x, t)) (fx (x0, t)) = psi (x, t)" @@ -2107,9 +2074,7 @@ lemma continuous_within_tendsto_compose': fixes f::"'a::t2_space \ 'b::topological_space" - assumes "continuous (at a within S) f" - "\n. x n \ S" - "(x \ a) F " + assumes "continuous (at a within S) f" "\n. x n \ S" "(x \ a) F " shows "((\n. f (x n)) \ f a) F" using always_eventually assms continuous_within_tendsto_compose by blast @@ -2119,7 +2084,7 @@ (\x. (\n::nat. x n \ S) \ (x \ a) sequentially \ ((f \ x) \ f a) sequentially)" using continuous_within_tendsto_compose'[of a S f _ sequentially] - continuous_within_sequentiallyI[of a S f] + using continuous_within_sequentiallyI[of a S f] by (auto simp: o_def) lemma continuous_at_sequentiallyI: @@ -2215,7 +2180,7 @@ shows "continuous_on A ((+) a \ g) = continuous_on A g" proof - have g: "g = (\x. -a + x) \ ((\x. a + x) \ g)" - by (rule ext) simp + by force show ?thesis by (metis (no_types, opaque_lifting) g continuous_on_compose homeomorphism_def homeomorphism_translation) qed @@ -2245,8 +2210,19 @@ "S homeomorphic T \ (\f g. (\x\S. f(x) \ T \ (g(f(x)) = x)) \ (\y\T. g(y) \ S \ (f(g(y)) = y)) \ - continuous_on S f \ continuous_on T g)" - by (smt (verit, ccfv_threshold) homeomorphic_def homeomorphismI homeomorphism_def image_eqI image_subset_iff) + continuous_on S f \ continuous_on T g)" (is "?L=?R") +proof + assume "S homeomorphic T" + then obtain f g where \
: "homeomorphism S T f g" + using homeomorphic_def by blast + show ?R + proof (intro exI conjI) + show "\x\S. f x \ T \ g (f x) = x" "\y\T. g y \ S \ f (g y) = y" + by (metis "\
" homeomorphism_def imageI)+ + show "continuous_on S f" "continuous_on T g" + using "\
" homeomorphism_def by blast+ + qed +qed (force simp: homeomorphic_def homeomorphism_def image_iff) lemma homeomorphicI [intro?]: "\f ` S = T; g ` T = S; @@ -2291,18 +2267,21 @@ lemma homeomorphic_finite: fixes S :: "'a::t1_space set" and T :: "'b::t1_space set" assumes "finite T" - shows "S homeomorphic T \ finite S \ finite T \ card S = card T" (is "?lhs = ?rhs") + shows "S homeomorphic T \ finite S \ card S = card T" (is "?lhs = ?rhs") proof assume "S homeomorphic T" with assms show ?rhs by (metis (full_types) card_image_le finite_imageI homeomorphic_def homeomorphism_def le_antisym) next assume R: ?rhs - with finite_same_card_bij obtain h where "bij_betw h S T" + with finite_same_card_bij assms obtain h where h: "bij_betw h S T" by auto - with R show ?lhs - apply (simp only: homeomorphic_def homeomorphism_def continuous_on_finite) - by (smt (verit, ccfv_SIG) bij_betw_imp_surj_on bij_betw_inv_into bij_betw_inv_into_left bij_betw_inv_into_right) + show ?lhs + unfolding homeomorphic_def homeomorphism_def + proof (intro exI conjI) + show "continuous_on S h" "continuous_on T (inv_into S h)" + by (simp_all add: assms R continuous_on_finite) + qed (use h in \auto simp: bij_betw_def\) qed text \Relatively weak hypotheses if a set is compact.\ diff -r 819688d4cc45 -r 62afd98e3f3e src/HOL/Analysis/Euclidean_Space.thy --- a/src/HOL/Analysis/Euclidean_Space.thy Wed Apr 16 21:13:33 2025 +0100 +++ b/src/HOL/Analysis/Euclidean_Space.thy Thu Apr 17 22:57:26 2025 +0100 @@ -58,10 +58,7 @@ by (simp add: inner_sum_left sum.If_cases inner_Basis) lemma (in euclidean_space) Basis_zero [simp]: "0 \ Basis" -proof - assume "0 \ Basis" thus "False" - using inner_Basis [of 0 0] by simp -qed + using local.inner_same_Basis by fastforce lemma (in euclidean_space) nonzero_Basis: "u \ Basis \ u \ 0" by clarsimp @@ -121,15 +118,17 @@ lemma (in euclidean_space) euclidean_representation_sum_fun: "(\x. \b\Basis. inner (f x) b *\<^sub>R b) = f" - by (rule ext) (simp add: euclidean_representation_sum) + by (force simp: euclidean_representation_sum) lemma euclidean_isCont: assumes "\b. b \ Basis \ isCont (\x. (inner (f x) b) *\<^sub>R b) x" - shows "isCont f x" - apply (subst euclidean_representation_sum_fun [symmetric]) - apply (rule isCont_sum) - apply (blast intro: assms) - done + shows "isCont f x" +proof - + have "isCont (\x. \b\Basis. inner (f x) b *\<^sub>R b) x" + by (simp add: assms) + then show ?thesis + by (simp add: euclidean_representation) +qed lemma DIM_positive [simp]: "0 < DIM('a::euclidean_space)" by (simp add: card_gt_0_iff) @@ -137,7 +136,6 @@ lemma DIM_ge_Suc0 [simp]: "Suc 0 \ card Basis" by (meson DIM_positive Suc_leI) - lemma sum_inner_Basis_scaleR [simp]: fixes f :: "'a::euclidean_space \ 'b::real_vector" assumes "b \ Basis" shows "(\i\Basis. (inner i b) *\<^sub>R f i) = f b" @@ -160,9 +158,9 @@ case False have "(\k\Basis. inner (if k = i then f i *\<^sub>R i else g k *\<^sub>R k) j) = (\k\Basis. if k = j then g k else 0)" - apply (rule sum.cong) - using False assms by (auto simp: inner_Basis) - also have "... = g j" + using False assms + by (intro sum.cong) (auto simp: inner_Basis) + also have "\ = g j" using assms by auto finally show ?thesis using False by (auto simp: inner_sum_left) @@ -182,10 +180,8 @@ by (metis Basis_le_norm le_less_trans) lemma norm_le_l1: "norm x \ (\b\Basis. \inner x b\)" - apply (subst euclidean_representation[of x, symmetric]) - apply (rule order_trans[OF norm_sum]) - apply (auto intro!: sum_mono) - done + by (metis (no_types, lifting) order.refl euclidean_representation mult.right_neutral + norm_Basis norm_scaleR sum_norm_le) lemma sum_norm_allsubsets_bound: fixes f :: "'a \ 'n::euclidean_space" @@ -364,14 +360,14 @@ "finite_dimensional_vector_space (*\<^sub>R) Basis" proof unfold_locales show "finite (Basis::'a set)" by (metis finite_Basis) - show "real_vector.independent (Basis::'a set)" - unfolding dependent_def dependent_raw_def[symmetric] - apply (subst span_finite) - apply simp - apply clarify + have "\a::'a. \u. \a \ Basis; a = (\v\Basis - {a}. u v *\<^sub>R v)\ \ False" apply (drule_tac f="inner a" in arg_cong) apply (simp add: inner_Basis inner_sum_right eq_commute) done + then + show "real_vector.independent (Basis::'a set)" + unfolding dependent_def dependent_raw_def[symmetric] + by (subst span_finite) auto show "module.span (*\<^sub>R) Basis = UNIV" unfolding span_finite [OF finite_Basis] span_raw_def[symmetric] by (auto intro!: euclidean_representation[symmetric]) diff -r 819688d4cc45 -r 62afd98e3f3e src/HOL/Analysis/Infinite_Sum.thy --- a/src/HOL/Analysis/Infinite_Sum.thy Wed Apr 16 21:13:33 2025 +0100 +++ b/src/HOL/Analysis/Infinite_Sum.thy Thu Apr 17 22:57:26 2025 +0100 @@ -300,15 +300,14 @@ \(\F. sum f (F \ A)) = sum f \ (\F. F \ A)\ by fastforce qed - with limB have "((\F. sum f F - sum f (F\A)) \ b - a) (finite_subsets_at_top B)" + with limB have \
: "((\F. sum f F - sum f (F\A)) \ b - a) (finite_subsets_at_top B)" using tendsto_diff by blast have "sum f X - sum f (X \ A) = sum f (X - A)" if "finite X" and "X \ B" for X :: "'a set" using that by (metis add_diff_cancel_left' sum.Int_Diff) hence "\\<^sub>F x in finite_subsets_at_top B. sum f x - sum f (x \ A) = sum f (x - A)" by (rule eventually_finite_subsets_at_top_weakI) hence "((\F. sum f (F-A)) \ b - a) (finite_subsets_at_top B)" - using tendsto_cong [THEN iffD1 , rotated] - \((\F. sum f F - sum f (F \ A)) \ b - a) (finite_subsets_at_top B)\ by fastforce + using \
tendsto_cong by fastforce hence "(sum f \ b - a) (filtermap (\F. F-A) (finite_subsets_at_top B))" by (subst tendsto_compose_filtermap[symmetric], simp add: o_def) thus ?thesis @@ -429,8 +428,7 @@ assumes \\x. x \ A \ f x \ g x\ assumes \x \ A\ \f x < g x\ shows "a < b" - by (rule has_sum_strict_mono_neutral[OF assms(1,2), where x = x]) - (use assms(3-) in auto) + using assms has_sum_strict_mono_neutral by force lemma has_sum_finite_approximation: fixes f :: "'a \ 'b::{comm_monoid_add,metric_space}" diff -r 819688d4cc45 -r 62afd98e3f3e src/HOL/Analysis/Inner_Product.thy --- a/src/HOL/Analysis/Inner_Product.thy Wed Apr 16 21:13:33 2025 +0100 +++ b/src/HOL/Analysis/Inner_Product.thy Thu Apr 17 22:57:26 2025 +0100 @@ -119,8 +119,6 @@ also have "\ = inner x x - (inner x y)\<^sup>2 / inner y y" by (simp add: power2_eq_square inner_commute) finally have "0 \ inner x x - (inner x y)\<^sup>2 / inner y y" . - hence "(inner x y)\<^sup>2 / inner y y \ inner x x" - by (simp add: le_diff_eq) thus "(inner x y)\<^sup>2 \ inner x x * inner y y" by (simp add: pos_divide_le_eq y) qed @@ -428,10 +426,7 @@ "\GDERIV f x :> df; GDERIV g x :> dg\ \ GDERIV (\x. f x * g x) x :> scaleR (f x) dg + scaleR (g x) df" unfolding gderiv_def - apply (rule has_derivative_subst) - apply (erule (1) has_derivative_mult) - apply (simp add: inner_add ac_simps) - done + by (auto simp: inner_add ac_simps intro: has_derivative_subst [OF has_derivative_mult]) lemma GDERIV_inverse: "\GDERIV f x :> df; f x \ 0\ diff -r 819688d4cc45 -r 62afd98e3f3e src/HOL/Analysis/L2_Norm.thy --- a/src/HOL/Analysis/L2_Norm.thy Wed Apr 16 21:13:33 2025 +0100 +++ b/src/HOL/Analysis/L2_Norm.thy Thu Apr 17 22:57:26 2025 +0100 @@ -59,10 +59,7 @@ lemma L2_set_right_distrib: "0 \ r \ r * L2_set f A = L2_set (\x. r * f x) A" unfolding L2_set_def - apply (simp add: power_mult_distrib) - apply (simp add: sum_distrib_left [symmetric]) - apply (simp add: real_sqrt_mult sum_nonneg) - done + by (simp add: power_mult_distrib real_sqrt_mult sum_nonneg flip: sum_distrib_left) lemma L2_set_left_distrib: "0 \ r \ L2_set f A * r = L2_set (\x. f x * r) A" diff -r 819688d4cc45 -r 62afd98e3f3e src/HOL/Analysis/Poly_Roots.thy --- a/src/HOL/Analysis/Poly_Roots.thy Wed Apr 16 21:13:33 2025 +0100 +++ b/src/HOL/Analysis/Poly_Roots.thy Thu Apr 17 22:57:26 2025 +0100 @@ -103,13 +103,10 @@ qed qed -lemma norm_lemma_xy: assumes "\b\ + 1 \ norm(y) - a" "norm(x) \ a" shows "b \ norm(x + y)" -proof - - have "b \ norm y - norm x" - using assms by linarith - then show ?thesis - by (metis (no_types) add.commute norm_diff_ineq order_trans) -qed +lemma norm_lemma_xy: + assumes "\b\ + 1 \ norm(y) - a" "norm(x) \ a" + shows "b \ norm(x + y)" + by (smt (verit) add.commute assms norm_diff_ineq) proposition polyfun_extremal: fixes c :: "nat \ 'a::real_normed_div_algebra" @@ -142,8 +139,7 @@ then have "\B\ * 2 + 2 \ norm z ^ (Suc n) * norm (c (Suc n))" by (metis \1 \ norm z\ order.trans mult_right_mono norm_ge_zero self_le_power zero_less_Suc) then show "B \ norm ((\i\n. c i * z^i) + c (Suc n) * (z * z ^ n))" using M les - apply auto - apply (rule norm_lemma_xy [where a = "norm (c (Suc n)) * norm z ^ (Suc n) / 2"]) + apply (intro norm_lemma_xy [where a = "norm (c (Suc n)) * norm z ^ (Suc n) / 2"]) apply (simp_all add: norm_mult norm_power) done qed @@ -173,11 +169,9 @@ then have extr_prem: "\ (\k\n. b k \ 0) \ \k. k \ 0 \ k \ Suc n \ c k \ 0" by (metis Suc.prems le0 minus_zero mult_zero_right) have "\k\n. b k \ 0" - apply (rule ccontr) - using polyfun_extremal [OF extr_prem, of 1] - apply (auto simp: eventually_at_infinity b simp del: sum.atMost_Suc) - apply (drule_tac x="of_real ba" in spec, simp) - done + using polyfun_extremal [OF extr_prem, of 1] + apply (simp add: eventually_at_infinity b del: sum.atMost_Suc) + by (metis norm_of_nat real_arch_simple) then show ?thesis using Suc.IH [of b] ins_ab by (auto simp: card_insert_if) qed simp @@ -220,11 +214,9 @@ fixes c :: "nat \ 'a::{comm_ring,real_normed_div_algebra}" shows "(\z. (\i\n. c i * z^i) = k) \ c 0 = k \ (\k. k \ 0 \ k \ n \ c k = 0)" proof - - {fix z - have "(\i\n. c i * z^i) = (\i\n. (if i = 0 then c 0 - k else c i) * z^i) + k" - by (induct n) auto - } then - have "(\z. (\i\n. c i * z^i) = k) \ (\z. (\i\n. (if i = 0 then c 0 - k else c i) * z^i) = 0)" + have "\z. (\i\n. c i * z^i) = (\i\n. (if i = 0 then c 0 - k else c i) * z^i) + k" + by (induct n) auto + then have "(\z. (\i\n. c i * z^i) = k) \ (\z. (\i\n. (if i = 0 then c 0 - k else c i) * z^i) = 0)" by auto also have "... \ c 0 = k \ (\k. k \ 0 \ k \ n \ c k = 0)" by (auto simp: polyfun_eq_0)