# HG changeset patch # User nipkow # Date 1545936521 -3600 # Node ID f9bf65d90b698687c596ee35180e53670194ec45 # Parent 04e54f57a8691dacb23d28760b17f50c6d4bf436# Parent 2a4c8a2a3f8e3fa056e386b24b6ff476be9d7ab0 merged diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Abstract_Topology.thy --- a/src/HOL/Analysis/Abstract_Topology.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Abstract_Topology.thy Thu Dec 27 19:48:41 2018 +0100 @@ -1979,7 +1979,7 @@ definition connected_space :: "'a topology \ bool" where "connected_space X \ - ~(\E1 E2. openin X E1 \ openin X E2 \ + \(\E1 E2. openin X E1 \ openin X E2 \ topspace X \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" definition connectedin :: "'a topology \ 'a set \ bool" where @@ -2093,10 +2093,10 @@ lemma connectedin_closedin: "connectedin X S \ S \ topspace X \ - ~(\E1 E2. closedin X E1 \ closedin X E2 \ + \(\E1 E2. closedin X E1 \ closedin X E2 \ S \ (E1 \ E2) \ (E1 \ E2 \ S = {}) \ - ~(E1 \ S = {}) \ ~(E2 \ S = {}))" + \(E1 \ S = {}) \ \(E2 \ S = {}))" proof - have *: "(\E1:: 'a set. \E2:: 'a set. (\T1:: 'a set. P1 T1 \ E1 = f1 T1) \ (\T2:: 'a set. P2 T2 \ E2 = f2 T2) \ R E1 E2) \ (\T1 T2. P1 T1 \ P2 T2 \ R(f1 T1) (f2 T2))" for P1 f1 P2 f2 R @@ -2269,7 +2269,7 @@ unfolding connectedin_eq_not_separated_subset by blast lemma connectedin_nonseparated_union: - "\connectedin X S; connectedin X T; ~separatedin X S T\ \ connectedin X (S \ T)" + "\connectedin X S; connectedin X T; \separatedin X S T\ \ connectedin X (S \ T)" apply (simp add: connectedin_eq_not_separated_subset, auto) apply (metis (no_types, hide_lams) Diff_subset_conv Diff_triv disjoint_iff_not_equal separatedin_mono sup_commute) apply (metis (no_types, hide_lams) Diff_subset_conv Diff_triv disjoint_iff_not_equal separatedin_mono separatedin_sym sup_commute) diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Brouwer_Fixpoint.thy --- a/src/HOL/Analysis/Brouwer_Fixpoint.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Brouwer_Fixpoint.thy Thu Dec 27 19:48:41 2018 +0100 @@ -1060,7 +1060,7 @@ text \More properties of ARs, ANRs and ENRs\ -lemma not_AR_empty [simp]: "~ AR({})" +lemma not_AR_empty [simp]: "\ AR({})" by (auto simp: AR_def) lemma ENR_empty [simp]: "ENR {}" @@ -3823,7 +3823,7 @@ lemma non_extensible_Borsuk_map: fixes a :: "'a :: euclidean_space" assumes "compact s" and cin: "c \ components(- s)" and boc: "bounded c" and "a \ c" - shows "~ (\g. continuous_on (s \ c) g \ + shows "\ (\g. continuous_on (s \ c) g \ g ` (s \ c) \ sphere 0 1 \ (\x \ s. g x = inverse(norm(x - a)) *\<^sub>R (x - a)))" proof - @@ -4534,7 +4534,7 @@ fixes a :: "'a :: euclidean_space" assumes "compact S" and "a \ S" shows "bounded (connected_component_set (- S) a) \ - ~(\c. homotopic_with (\x. True) S (sphere 0 1) + \(\c. homotopic_with (\x. True) S (sphere 0 1) (\x. inverse(norm(x - a)) *\<^sub>R (x - a)) (\x. c))" (is "?lhs = ?rhs") proof (cases "S = {}") @@ -4550,9 +4550,9 @@ by (simp add: \a \ S\) obtain r where "r>0" and r: "S \ ball 0 r" using bounded_subset_ballD \bounded S\ by blast - have "~ ?rhs \ ~ ?lhs" + have "\ ?rhs \ \ ?lhs" proof - assume notr: "~ ?rhs" + assume notr: "\ ?rhs" have nog: "\g. continuous_on (S \ connected_component_set (- S) a) g \ g ` (S \ connected_component_set (- S) a) \ sphere 0 1 \ (\x\S. g x = (x - a) /\<^sub>R norm (x - a))" @@ -4564,12 +4564,12 @@ using notr by (auto simp: nullhomotopic_into_sphere_extension [OF \closed S\ continuous_on_Borsuk_map [OF \a \ S\] False s01]) - with \a \ S\ show "~ ?lhs" + with \a \ S\ show "\ ?lhs" apply (clarsimp simp: Borsuk_map_into_sphere [of a S, symmetric] dest!: nog) apply (drule_tac x=g in spec) using continuous_on_subset by fastforce next - assume "~ ?lhs" + assume "\ ?lhs" then obtain b where b: "b \ connected_component_set (- S) a" and "r \ norm b" using bounded_iff linear by blast then have bnot: "b \ ball 0 r" @@ -4593,7 +4593,7 @@ ultimately have "homotopic_with (\x. True) S (sphere 0 1) (\x. (x - a) /\<^sub>R norm (x - a)) (\x. c)" by (meson homotopic_with_subset_left homotopic_with_trans r) - then show "~ ?rhs" + then show "\ ?rhs" by blast qed then show ?thesis by blast diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Cartesian_Euclidean_Space.thy --- a/src/HOL/Analysis/Cartesian_Euclidean_Space.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Cartesian_Euclidean_Space.thy Thu Dec 27 19:48:41 2018 +0100 @@ -651,7 +651,7 @@ lemma%unimportant Lim_component_eq_cart: fixes f :: "'a \ real^'n" - assumes net: "(f \ l) net" "~(trivial_limit net)" and ev:"eventually (\x. f(x)$i = b) net" + assumes net: "(f \ l) net" "\ trivial_limit net" and ev:"eventually (\x. f(x)$i = b) net" shows "l$i = b" using ev[unfolded order_eq_iff eventually_conj_iff] and Lim_component_ge_cart[OF net, of b i] and @@ -1060,7 +1060,7 @@ lemma%unimportant matrix_nonfull_linear_equations_eq: fixes A :: "real^'n^'m" - shows "(\x. (x \ 0) \ A *v x = 0) \ ~(rank A = CARD('n))" + shows "(\x. (x \ 0) \ A *v x = 0) \ rank A \ CARD('n)" by (meson matrix_left_invertible_injective full_rank_injective matrix_left_invertible_ker) lemma%unimportant rank_eq_0: "rank A = 0 \ A = 0" and rank_0 [simp]: "rank (0::real^'n^'m) = 0" diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Cauchy_Integral_Theorem.thy --- a/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Thu Dec 27 19:48:41 2018 +0100 @@ -941,7 +941,7 @@ case True then show ?thesis by (simp add: assms contour_integral_unique has_contour_integral_integral has_contour_integral_reversepath) next - case False then have "~ f contour_integrable_on (reversepath g)" + case False then have "\ f contour_integrable_on (reversepath g)" by (simp add: assms contour_integrable_reversepath_eq) with False show ?thesis by (simp add: not_integrable_contour_integral) qed @@ -5518,7 +5518,7 @@ and ul_f: "uniform_limit (path_image \) f l F" and noleB: "\t. t \ {0..1} \ norm (vector_derivative \ (at t)) \ B" and \: "valid_path \" - and [simp]: "~ (trivial_limit F)" + and [simp]: "\ trivial_limit F" shows "l contour_integrable_on \" "((\n. contour_integral \ (f n)) \ contour_integral \ l) F" proof - have "0 \ B" by (meson noleB [of 0] atLeastAtMost_iff norm_ge_zero order_refl order_trans zero_le_one) @@ -5586,7 +5586,7 @@ corollary%unimportant contour_integral_uniform_limit_circlepath: assumes "\\<^sub>F n::'a in F. (f n) contour_integrable_on (circlepath z r)" and "uniform_limit (sphere z r) f l F" - and "~ (trivial_limit F)" "0 < r" + and "\ trivial_limit F" "0 < r" shows "l contour_integrable_on (circlepath z r)" "((\n. contour_integral (circlepath z r) (f n)) \ contour_integral (circlepath z r) l) F" using assms by (auto simp: vector_derivative_circlepath norm_mult intro!: contour_integral_uniform_limit) @@ -6550,7 +6550,7 @@ lemma holomorphic_uniform_limit: assumes cont: "eventually (\n. continuous_on (cball z r) (f n) \ (f n) holomorphic_on ball z r) F" and ulim: "uniform_limit (cball z r) f g F" - and F: "~ trivial_limit F" + and F: "\ trivial_limit F" obtains "continuous_on (cball z r) g" "g holomorphic_on ball z r" proof (cases r "0::real" rule: linorder_cases) case less then show ?thesis by (force simp: ball_empty less_imp_le continuous_on_def holomorphic_on_def intro: that) @@ -6620,7 +6620,7 @@ assumes cont: "eventually (\n. continuous_on (cball z r) (f n) \ (\w \ ball z r. ((f n) has_field_derivative (f' n w)) (at w))) F" and ulim: "uniform_limit (cball z r) f g F" - and F: "~ trivial_limit F" and "0 < r" + and F: "\ trivial_limit F" and "0 < r" obtains g' where "continuous_on (cball z r) g" "\w. w \ ball z r \ (g has_field_derivative (g' w)) (at w) \ ((\n. f' n w) \ g' w) F" diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Complex_Analysis_Basics.thy --- a/src/HOL/Analysis/Complex_Analysis_Basics.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Complex_Analysis_Basics.thy Thu Dec 27 19:48:41 2018 +0100 @@ -212,7 +212,7 @@ lemma real_lim: fixes l::complex - assumes "(f \ l) F" and "~(trivial_limit F)" and "eventually P F" and "\a. P a \ f a \ \" + assumes "(f \ l) F" and "\ trivial_limit F" and "eventually P F" and "\a. P a \ f a \ \" shows "l \ \" proof (rule Lim_in_closed_set[OF closed_complex_Reals _ assms(2,1)]) show "eventually (\x. f x \ \) F" diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Complex_Transcendental.thy --- a/src/HOL/Analysis/Complex_Transcendental.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Complex_Transcendental.thy Thu Dec 27 19:48:41 2018 +0100 @@ -1133,21 +1133,21 @@ lemma cnj_tan: "cnj(tan z) = tan(cnj z)" by (simp add: cnj_cos cnj_sin tan_def) -lemma field_differentiable_at_tan: "~(cos z = 0) \ tan field_differentiable at z" +lemma field_differentiable_at_tan: "cos z \ 0 \ tan field_differentiable at z" unfolding field_differentiable_def using DERIV_tan by blast -lemma field_differentiable_within_tan: "~(cos z = 0) +lemma field_differentiable_within_tan: "cos z \ 0 \ tan field_differentiable (at z within s)" using field_differentiable_at_tan field_differentiable_at_within by blast -lemma continuous_within_tan: "~(cos z = 0) \ continuous (at z within s) tan" +lemma continuous_within_tan: "cos z \ 0 \ continuous (at z within s) tan" using continuous_at_imp_continuous_within isCont_tan by blast -lemma continuous_on_tan [continuous_intros]: "(\z. z \ s \ ~(cos z = 0)) \ continuous_on s tan" +lemma continuous_on_tan [continuous_intros]: "(\z. z \ s \ cos z \ 0) \ continuous_on s tan" by (simp add: continuous_at_imp_continuous_on) -lemma holomorphic_on_tan: "(\z. z \ s \ ~(cos z = 0)) \ tan holomorphic_on s" +lemma holomorphic_on_tan: "(\z. z \ s \ cos z \ 0) \ tan holomorphic_on s" by (simp add: field_differentiable_within_tan holomorphic_on_def) @@ -1670,7 +1670,7 @@ lemma Ln_minus: assumes "z \ 0" - shows "Ln(-z) = (if Im(z) \ 0 \ ~(Re(z) < 0 \ Im(z) = 0) + shows "Ln(-z) = (if Im(z) \ 0 \ \(Re(z) < 0 \ Im(z) = 0) then Ln(z) + \ * pi else Ln(z) - \ * pi)" (is "_ = ?rhs") using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Conformal_Mappings.thy --- a/src/HOL/Analysis/Conformal_Mappings.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Conformal_Mappings.thy Thu Dec 27 19:48:41 2018 +0100 @@ -359,7 +359,7 @@ assumes holf: "f holomorphic_on S" and S: "open S" and "connected S" and "open U" and "U \ S" - and fne: "~ f constant_on S" + and fne: "\ f constant_on S" shows "open (f ` U)" proof - have *: "open (f ` U)" @@ -468,7 +468,7 @@ assumes holf: "f holomorphic_on S" and S: "open S" and "open U" "U \ S" - and fnc: "\X. \open X; X \ S; X \ {}\ \ ~ f constant_on X" + and fnc: "\X. \open X; X \ S; X \ {}\ \ \ f constant_on X" shows "open (f ` U)" proof - have "S = \(components S)" by simp @@ -518,7 +518,7 @@ assume "\ f constant_on S" then have "open (f ` U)" using open_mapping_thm assms by blast - moreover have "~ open (f ` U)" + moreover have "\ open (f ` U)" proof - have "\t. cmod (f \ - t) < e \ t \ f ` U" if "0 < e" for e apply (rule_tac x="if 0 < Re(f \) then f \ + (e/2) else f \ - (e/2)" in exI) @@ -735,7 +735,7 @@ lemma holomorphic_factor_zero_nonconstant: assumes holf: "f holomorphic_on S" and S: "open S" "connected S" and "\ \ S" "f \ = 0" - and nonconst: "~ f constant_on S" + and nonconst: "\ f constant_on S" obtains g r n where "0 < n" "0 < r" "ball \ r \ S" "g holomorphic_on ball \ r" @@ -1156,7 +1156,7 @@ obtain r where "r > 0" "ball \ r \ S" "inj_on f (ball \ r)" by (blast intro: that has_complex_derivative_locally_injective [OF assms]) then have \: "\ \ ball \ r" by simp - then have nc: "~ f constant_on ball \ r" + then have nc: "\ f constant_on ball \ r" using \inj_on f (ball \ r)\ injective_not_constant by fastforce have holf': "f holomorphic_on ball \ r" using \ball \ r \ S\ holf holomorphic_on_subset by blast @@ -2132,7 +2132,7 @@ case False then have "t > 0" using 2 by (force simp: zero_less_mult_iff) - have "~ ball a t \ s \ ball a t \ frontier s \ {}" + have "\ ball a t \ s \ ball a t \ frontier s \ {}" apply (rule connected_Int_frontier [of "ball a t" s], simp_all) using \0 < t\ \a \ s\ centre_in_ball apply blast done diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Connected.thy --- a/src/HOL/Analysis/Connected.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Connected.thy Thu Dec 27 19:48:41 2018 +0100 @@ -3280,7 +3280,7 @@ lemma continuous_on_cases_local: "\closedin (subtopology euclidean (s \ t)) s; closedin (subtopology euclidean (s \ t)) t; continuous_on s f; continuous_on t g; - \x. \x \ s \ ~P x \ x \ t \ P x\ \ f x = g x\ + \x. \x \ s \ \P x \ x \ t \ P x\ \ f x = g x\ \ continuous_on (s \ t) (\x. if P x then f x else g x)" by (rule continuous_on_Un_local) (auto intro: continuous_on_eq) @@ -4983,7 +4983,7 @@ shows "connected (S \ T)" proof - have *: "\\x y. P x y \ P y x; \x y. P x y \ S \ x \ S \ y; - \x y. \P x y; S \ x\ \ False\ \ ~(\x y. (P x y))" for P + \x y. \P x y; S \ x\ \ False\ \ \(\x y. (P x y))" for P by metis show ?thesis unfolding connected_closedin_eq diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Continuous_Extension.thy --- a/src/HOL/Analysis/Continuous_Extension.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Continuous_Extension.thy Thu Dec 27 19:48:41 2018 +0100 @@ -2,7 +2,7 @@ Authors: LC Paulson, based on material from HOL Light *) -section \Continuous extensions of functions: Urysohn's lemma, Dugundji extension theorem, Tietze\ +section \Continuous Extensions of Functions\ theory Continuous_Extension imports Starlike @@ -112,7 +112,9 @@ qed -subsection\Urysohn's lemma (for Euclidean spaces, where the proof is easy using distances)\ +subsection\Urysohn's Lemma for Euclidean Spaces\ + +text \For Euclidean spaces the proof is easy using distances.\ lemma Urysohn_both_ne: assumes US: "closedin (subtopology euclidean U) S" @@ -130,7 +132,7 @@ using \T \ {}\ UT setdist_eq_0_closedin by auto have sdpos: "0 < setdist {x} S + setdist {x} T" if "x \ U" for x proof - - have "~ (setdist {x} S = 0 \ setdist {x} T = 0)" + have "\ (setdist {x} S = 0 \ setdist {x} T = 0)" using assms by (metis IntI empty_iff setdist_eq_0_closedin that) then show ?thesis by (metis add.left_neutral add.right_neutral add_pos_pos linorder_neqE_linordered_idom not_le setdist_pos_le) @@ -295,7 +297,7 @@ using assms by (auto intro: Urysohn_local [of UNIV S T a b]) -subsection\The Dugundji extension theorem and Tietze variants as corollaries\ +subsection\The Dugundji Extension Theorem and Tietze Variants\ text%important\J. Dugundji. An extension of Tietze's theorem. Pacific J. Math. Volume 1, Number 3 (1951), 353-367. https://projecteuclid.org/euclid.pjm/1103052106\ diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Analysis/Convex_Euclidean_Space.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy Thu Dec 27 19:48:41 2018 +0100 @@ -1343,7 +1343,9 @@ by (simp add: affine_def algebra_simps) (metis distrib_right mult.left_neutral) -subsubsection%unimportant \Some explicit formulations (from Lars Schewe)\ +subsubsection%unimportant \Some explicit formulations\ + +text "Formalized by Lars Schewe." lemma affine: fixes V::"'a::real_vector set" @@ -2070,7 +2072,9 @@ qed -subsection \Affine dependence and consequential theorems (from Lars Schewe)\ +subsection \Affine dependence and consequential theorems\ + +text "Formalized by Lars Schewe." definition%important affine_dependent :: "'a::real_vector set \ bool" where "affine_dependent s \ (\x\s. x \ affine hull (s - {x}))" @@ -2082,11 +2086,11 @@ done lemma affine_independent_subset: - shows "\~ affine_dependent t; s \ t\ \ ~ affine_dependent s" + shows "\\ affine_dependent t; s \ t\ \ \ affine_dependent s" by (metis affine_dependent_subset) lemma affine_independent_Diff: - "~ affine_dependent s \ ~ affine_dependent(s - t)" + "\ affine_dependent s \ \ affine_dependent(s - t)" by (meson Diff_subset affine_dependent_subset) proposition affine_dependent_explicit: @@ -2601,7 +2605,9 @@ qed (use assms in \auto simp: convex_explicit\) -subsubsection%unimportant \Another formulation from Lars Schewe\ +subsubsection%unimportant \Another formulation\ + +text "Formalized by Lars Schewe." lemma convex_hull_explicit: fixes p :: "'a::real_vector set" @@ -3427,7 +3433,7 @@ lemma affine_independent_iff_card: fixes s :: "'a::euclidean_space set" - shows "~ affine_dependent s \ finite s \ aff_dim s = int(card s) - 1" + shows "\ affine_dependent s \ finite s \ aff_dim s = int(card s) - 1" apply (rule iffI) apply (simp add: aff_dim_affine_independent aff_independent_finite) by (metis affine_basis_exists [of s] aff_dim_unique card_subset_eq diff_add_cancel of_nat_eq_iff) @@ -3652,7 +3658,7 @@ lemma affine_independent_card_dim_diffs: fixes S :: "'a :: euclidean_space set" - assumes "~ affine_dependent S" "a \ S" + assumes "\ affine_dependent S" "a \ S" shows "card S = dim {x - a|x. x \ S} + 1" proof - have 1: "{b - a|b. b \ (S - {a})} \ {x - a|x. x \ S}" by auto @@ -3766,7 +3772,7 @@ show ?thesis proof safe assume 0: "aff_dim S = 0" - have "~ {a,b} \ S" if "b \ a" for b + have "\ {a,b} \ S" if "b \ a" for b by (metis "0" aff_dim_2 aff_dim_subset not_one_le_zero that) then show "\a. S = {a}" using \a \ S\ by blast @@ -3798,7 +3804,7 @@ lemma disjoint_affine_hull: fixes s :: "'n::euclidean_space set" - assumes "~ affine_dependent s" "t \ s" "u \ s" "t \ u = {}" + assumes "\ affine_dependent s" "t \ s" "u \ s" "t \ u = {}" shows "(affine hull t) \ (affine hull u) = {}" proof - have "finite s" using assms by (simp add: aff_independent_finite) @@ -3813,7 +3819,7 @@ have [simp]: "s \ t = t" "s \ - t \ u = u" using assms by auto have "sum c s = 0" by (simp add: c_def comm_monoid_add_class.sum.If_cases \finite s\ sum_negf) - moreover have "~ (\v\s. c v = 0)" + moreover have "\ (\v\s. c v = 0)" by (metis (no_types) IntD1 \s \ t = t\ a1 c_def sum_not_0 zero_neq_one) moreover have "(\v\s. c v *\<^sub>R v) = 0" by (simp add: c_def if_smult sum_negf @@ -5747,7 +5753,9 @@ qed auto -subsection \Radon's theorem (from Lars Schewe)\ +subsection \Radon's theorem\ + +text "Formalized by Lars Schewe." lemma radon_ex_lemma: assumes "finite c" "affine_dependent c" diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Cross3.thy --- a/src/HOL/Analysis/Cross3.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Cross3.thy Thu Dec 27 19:48:41 2018 +0100 @@ -99,7 +99,7 @@ by (force simp add: axis_def cross3_simps)+ lemma cross_basis_nonzero: - "u \ 0 \ ~(u \ axis 1 1 = 0) \ ~(u \ axis 2 1 = 0) \ ~(u \ axis 3 1 = 0)" + "u \ 0 \ u \ axis 1 1 \ 0 \ u \ axis 2 1 \ 0 \ u \ axis 3 1 \ 0" by (clarsimp simp add: axis_def cross3_simps) (metis vector_3 exhaust_3) lemma cross_dot_cancel: diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy --- a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Thu Dec 27 19:48:41 2018 +0100 @@ -1195,8 +1195,8 @@ proof (rule starlike_negligible [OF \closed S\, of a]) fix c x assume cx: "a + c *\<^sub>R x \ S" "0 \ c" "a + x \ S" - with star have "~ (c < 1)" by auto - moreover have "~ (c > 1)" + with star have "\ (c < 1)" by auto + moreover have "\ (c > 1)" using star [of "1/c" "c *\<^sub>R x"] cx by force ultimately show "c = 1" by arith qed diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Further_Topology.thy --- a/src/HOL/Analysis/Further_Topology.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Further_Topology.thy Thu Dec 27 19:48:41 2018 +0100 @@ -366,7 +366,7 @@ by (meson f nullhomotopic_from_contractible contractible_sphere that) next case False - with \~ s \ 0\ have "r > 0" "s > 0" by auto + with \\ s \ 0\ have "r > 0" "s > 0" by auto show ?thesis apply (rule inessential_spheremap_lowdim_gen [of "cball a r" "cball b s" f]) using \0 < r\ \0 < s\ assms(1) @@ -414,7 +414,7 @@ assumes fin: "finite \" and "\S. S \ \ \ \g. continuous_on S g \ g ` S \ T \ (\x \ S \ K. g x = h x)" and "\S. S \ \ \ closed S" - and K: "\X Y. \X \ \; Y \ \; ~ X \ Y; ~ Y \ X\ \ X \ Y \ K" + and K: "\X Y. \X \ \; Y \ \; \ X \ Y; \ Y \ X\ \ X \ Y \ K" shows "\g. continuous_on (\\) g \ g ` (\\) \ T \ (\x \ \\ \ K. g x = h x)" apply (simp add: Union_maximal_sets [OF fin, symmetric]) apply (rule extending_maps_Union_aux) @@ -567,7 +567,7 @@ have clo: "closed S" if "S \ \ \ ?Faces" for S using that \\ \ \\ face_of_polytope_polytope poly polytope_imp_closed by blast have K: "X \ Y \ \(\ \ {D. \C\\. D face_of C \ aff_dim D < int p})" - if "X \ \ \ ?Faces" "Y \ \ \ ?Faces" "~ Y \ X" for X Y + if "X \ \ \ ?Faces" "Y \ \ \ ?Faces" "\ Y \ X" for X Y proof - have ff: "X \ Y face_of X \ X \ Y face_of Y" if XY: "X face_of D" "Y face_of E" and DE: "D \ \" "E \ \" for D E @@ -669,7 +669,7 @@ assumes "finite \" and \: "\X. X \ \ \ \a g. a \ U \ continuous_on (X - {a}) g \ g ` (X - {a}) \ T \ (\x \ X \ K. g x = h x)" and clo: "\X. X \ \ \ closed X" - and K: "\X Y. \X \ \; Y \ \; ~(X \ Y); ~(Y \ X)\ \ X \ Y \ K" + and K: "\X Y. \X \ \; Y \ \; \ X \ Y; \ Y \ X\ \ X \ Y \ K" obtains C g where "finite C" "disjnt C U" "card C \ card \" "continuous_on (\\ - C) g" "g ` (\\ - C) \ T" "\x. x \ (\\ - C) \ K \ g x = h x" @@ -1188,8 +1188,8 @@ then have "ball a d \ U \ C" by auto obtain h k where homhk: "homeomorphism (S \ C) (S \ C) h k" - and subC: "{x. (~ (h x = x \ k x = x))} \ C" - and bou: "bounded {x. (~ (h x = x \ k x = x))}" + and subC: "{x. (\ (h x = x \ k x = x))} \ C" + and bou: "bounded {x. (\ (h x = x \ k x = x))}" and hin: "\x. x \ C \ K \ h x \ ball a d \ U" proof (rule homeomorphism_grouping_points_exists_gen [of C "ball a d \ U" "C \ K" "S \ C"]) show "openin (subtopology euclidean C) (ball a d \ U)" @@ -1539,7 +1539,7 @@ by (simp add: assms compact_imp_bounded) then obtain b where b: "S \ cbox (-b) b" using bounded_subset_cbox_symmetric by blast - define LU where "LU \ L \ (\ {C \ components (T - S). ~bounded C} - cbox (-(b+One)) (b+One))" + define LU where "LU \ L \ (\ {C \ components (T - S). \bounded C} - cbox (-(b+One)) (b+One))" obtain K g where "finite K" "K \ LU" "K \ T" "disjnt K S" and contg: "continuous_on (T - K) g" and gim: "g ` (T - K) \ rel_frontier U" @@ -1699,7 +1699,7 @@ theorem%important Borsuk_separation_theorem_gen: fixes S :: "'a::euclidean_space set" assumes "compact S" - shows "(\c \ components(- S). ~bounded c) \ + shows "(\c \ components(- S). \bounded c) \ (\f. continuous_on S f \ f ` S \ sphere (0::'a) 1 \ (\c. homotopic_with (\x. True) S (sphere 0 1) f (\x. c)))" (is "?lhs = ?rhs") @@ -5204,7 +5204,7 @@ proof (rule ccontr) let ?S = "S \ \{C \ components(- (S \ T)). frontier C \ S}" let ?T = "T \ \{C \ components(- (S \ T)). frontier C \ T}" - assume "~ thesis" + assume "\ thesis" with that have *: "frontier C \ S = {} \ frontier C \ T = {}" if C: "C \ components (- (S \ T))" "C \ {}" for C using C by blast diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Great_Picard.thy --- a/src/HOL/Analysis/Great_Picard.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Great_Picard.thy Thu Dec 27 19:48:41 2018 +0100 @@ -874,7 +874,7 @@ have le_er: "cmod (\ n x / (x - y) - \ n x / (x - z)) \ e / r" if "cmod (x - z) = r/3 + r/3" for x proof - - have "~ (cmod (x - y) < r/3)" + have "\ (cmod (x - y) < r/3)" using y_near_z(1) that \M > 0\ \r > 0\ by (metis (full_types) norm_diff_triangle_less norm_minus_commute order_less_irrefl) then have r4_le_xy: "r/4 \ cmod (x - y)" @@ -1000,7 +1000,7 @@ and holf: "\n::nat. \ n holomorphic_on S" and holg: "g holomorphic_on S" and ul_g: "\K. \compact K; K \ S\ \ uniform_limit K \ g sequentially" - and nonconst: "~ g constant_on S" + and nonconst: "\ g constant_on S" and nz: "\n z. z \ S \ \ n z \ 0" and "z0 \ S" shows "g z0 \ 0" @@ -1164,7 +1164,7 @@ and holf: "\n::nat. \ n holomorphic_on S" and holg: "g holomorphic_on S" and ul_g: "\K. \compact K; K \ S\ \ uniform_limit K \ g sequentially" - and nonconst: "~ g constant_on S" + and nonconst: "\ g constant_on S" and inj: "\n. inj_on (\ n) S" shows "inj_on g S" proof%unimportant - @@ -1309,7 +1309,7 @@ unfolding open_subopen [of U] by (auto simp: U_def) fix v assume v: "v islimpt U" "v \ S" - have "~ (\r>0. \h\Y. r < cmod (h v))" + have "\ (\r>0. \h\Y. r < cmod (h v))" proof assume "\r>0. \h\Y. r < cmod (h v)" then have "\n. \h\Y. Suc n < cmod (h v)" diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Henstock_Kurzweil_Integration.thy --- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Thu Dec 27 19:48:41 2018 +0100 @@ -79,7 +79,7 @@ lemma content_pos_le [iff]: "0 \ content X" by simp -corollary content_nonneg [simp]: "~ content (cbox a b) < 0" +corollary content_nonneg [simp]: "\ content (cbox a b) < 0" using not_le by blast lemma content_pos_lt: "\i\Basis. a\i < b\i \ 0 < content (cbox a b)" @@ -287,12 +287,12 @@ definition integrable_on (infixr "integrable'_on" 46) where "f integrable_on i \ (\y. (f has_integral y) i)" -definition "integral i f = (SOME y. (f has_integral y) i \ ~ f integrable_on i \ y=0)" +definition "integral i f = (SOME y. (f has_integral y) i \ \ f integrable_on i \ y=0)" lemma integrable_integral[intro]: "f integrable_on i \ (f has_integral (integral i f)) i" unfolding integrable_on_def integral_def by (metis (mono_tags, lifting) someI_ex) -lemma not_integrable_integral: "~ f integrable_on i \ integral i f = 0" +lemma not_integrable_integral: "\ f integrable_on i \ integral i f = 0" unfolding integrable_on_def integral_def by blast lemma has_integral_integrable[dest]: "(f has_integral i) s \ f integrable_on s" @@ -356,7 +356,7 @@ lemma has_integral_iff: "(f has_integral i) S \ (f integrable_on S \ integral S f = i)" by blast -lemma eq_integralD: "integral k f = y \ (f has_integral y) k \ ~ f integrable_on k \ y=0" +lemma eq_integralD: "integral k f = y \ (f has_integral y) k \ \ f integrable_on k \ y=0" unfolding integral_def integrable_on_def apply (erule subst) apply (rule someI_ex) @@ -498,7 +498,7 @@ case True then show ?thesis by (force intro: has_integral_mult_left) next - case False then have "~ (\x. f x * c) integrable_on S" + case False then have "\ (\x. f x * c) integrable_on S" using has_integral_mult_left [of "(\x. f x * c)" _ S "inverse c"] by (auto simp add: mult.assoc) with False show ?thesis by (simp add: not_integrable_integral) @@ -615,7 +615,7 @@ case True with has_integral_cmul integrable_integral show ?thesis by fastforce next - case False then have "~ (\x. c *\<^sub>R f x) integrable_on S" + case False then have "\ (\x. c *\<^sub>R f x) integrable_on S" using has_integral_cmul [of "(\x. c *\<^sub>R f x)" _ S "inverse c"] by auto with False show ?thesis by (simp add: not_integrable_integral) qed @@ -630,7 +630,7 @@ case True then show ?thesis by (simp add: has_integral_neg integrable_integral integral_unique) next - case False then have "~ (\x. - f x) integrable_on S" + case False then have "\ (\x. - f x) integrable_on S" using has_integral_neg [of "(\x. - f x)" _ S ] by auto with False show ?thesis by (simp add: not_integrable_integral) qed diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Homeomorphism.thy --- a/src/HOL/Analysis/Homeomorphism.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Homeomorphism.thy Thu Dec 27 19:48:41 2018 +0100 @@ -201,7 +201,7 @@ lemma%unimportant segment_to_rel_frontier: fixes x :: "'a::euclidean_space" assumes S: "convex S" "bounded S" and x: "x \ rel_interior S" - and y: "y \ S" and xy: "~(x = y \ S = {x})" + and y: "y \ S" and xy: "\(x = y \ S = {x})" obtains z where "z \ rel_frontier S" "y \ closed_segment x z" "open_segment x z \ rel_interior S" proof (cases "x=y") diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Improper_Integral.thy --- a/src/HOL/Analysis/Improper_Integral.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Improper_Integral.thy Thu Dec 27 19:48:41 2018 +0100 @@ -924,7 +924,7 @@ moreover have "norm (\(x,K) \ T'. ?CI K h x - ?CI K f x) \ 2*\/3" proof - - define T'' where "T'' \ {(x,K) \ T'. ~ (K \ {x. x \ i \ c})}" + define T'' where "T'' \ {(x,K) \ T'. \ (K \ {x. x \ i \ c})}" then have "T'' \ T'" by auto then have "finite T''" diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Jordan_Curve.thy --- a/src/HOL/Analysis/Jordan_Curve.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Jordan_Curve.thy Thu Dec 27 19:48:41 2018 +0100 @@ -374,7 +374,7 @@ qed then have "connected_component (- (path_image u \ path_image d \ path_image g)) x y" by (simp add: Un_ac) - moreover have "~(connected_component (- (path_image c)) x y)" + moreover have "\(connected_component (- (path_image c)) x y)" by (metis (no_types, lifting) \\ bounded outer\ \bounded inner\ \x \ inner\ \y \ outer\ componentsE connected_component_eq inner mem_Collect_eq outer) ultimately show False by (auto simp: ud_Un [symmetric] connected_component_def) diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Ordered_Euclidean_Space.thy --- a/src/HOL/Analysis/Ordered_Euclidean_Space.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Ordered_Euclidean_Space.thy Thu Dec 27 19:48:41 2018 +0100 @@ -145,7 +145,7 @@ qed lemma%unimportant (in order) atLeastatMost_empty'[simp]: - "(~ a <= b) \ {a..b} = {}" + "(\ a \ b) \ {a..b} = {}" by (auto) instance real :: ordered_euclidean_space diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Path_Connected.thy --- a/src/HOL/Analysis/Path_Connected.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Path_Connected.thy Thu Dec 27 19:48:41 2018 +0100 @@ -591,7 +591,7 @@ using \d1 > 0\ \d2 > 0\ by (simp add: min_def) have dist2: "dist (1 / 2 + min (1 / 2) (min d1 d2) / 4) (1 / 2) < d2" using \d1 > 0\ \d2 > 0\ by (simp add: min_def dist_norm) - have [simp]: "~ min (1 / 2) (min d1 d2) \ 0" + have [simp]: "\ min (1 / 2) (min d1 d2) \ 0" using \d1 > 0\ \d2 > 0\ by (simp add: min_def) have "dist (g2 (min (1 / 2) (min d1 d2) / 2)) (g1 1) < e/2" "dist (g2 (min (1 / 2) (min d1 d2) / 2)) (g2 0) < e/2" @@ -670,7 +670,7 @@ have False if "g1 0 = g2 u" "0 \ u" "u \ 1" for u using * [of 0 "(u + 1) / 2"] that assms arc_distinct_ends [OF \?lhs\] by (auto simp: joinpaths_def pathstart_def pathfinish_def split_ifs divide_simps) - then have n1: "~ (pathstart g1 \ path_image g2)" + then have n1: "pathstart g1 \ path_image g2" unfolding pathstart_def path_image_def using atLeastAtMost_iff by blast show ?rhs using \?lhs\ @@ -2442,7 +2442,7 @@ lemma cobounded_unbounded_component: fixes s :: "'a :: euclidean_space set" assumes "bounded (-s)" - shows "\x. x \ s \ ~ bounded (connected_component_set s x)" + shows "\x. x \ s \ \ bounded (connected_component_set s x)" proof - obtain i::'a where i: "i \ Basis" using nonempty_Basis by blast @@ -2450,7 +2450,7 @@ using bounded_subset_ballD [OF assms, of 0] by auto then have *: "\x. B \ norm x \ x \ s" by (force simp: ball_def dist_norm) - have unbounded_inner: "~ bounded {x. inner i x \ B}" + have unbounded_inner: "\ bounded {x. inner i x \ B}" apply (auto simp: bounded_def dist_norm) apply (rule_tac x="x + (max B e + 1 + \i \ x\) *\<^sub>R i" in exI) apply simp @@ -2475,8 +2475,8 @@ lemma cobounded_unique_unbounded_component: fixes s :: "'a :: euclidean_space set" assumes bs: "bounded (-s)" and "2 \ DIM('a)" - and bo: "~ bounded(connected_component_set s x)" - "~ bounded(connected_component_set s y)" + and bo: "\ bounded(connected_component_set s x)" + "\ bounded(connected_component_set s y)" shows "connected_component_set s x = connected_component_set s y" proof - obtain i::'a where i: "i \ Basis" @@ -2507,7 +2507,7 @@ lemma cobounded_unbounded_components: fixes s :: "'a :: euclidean_space set" - shows "bounded (-s) \ \c. c \ components s \ ~bounded c" + shows "bounded (-s) \ \c. c \ components s \ \bounded c" by (metis cobounded_unbounded_component components_def imageI) lemma cobounded_unique_unbounded_components: @@ -2532,9 +2532,9 @@ "inside S \ {x. (x \ S) \ bounded(connected_component_set ( - S) x)}" definition%important outside where - "outside S \ -S \ {x. ~ bounded(connected_component_set (- S) x)}" - -lemma outside: "outside S = {x. ~ bounded(connected_component_set (- S) x)}" + "outside S \ -S \ {x. \ bounded(connected_component_set (- S) x)}" + +lemma outside: "outside S = {x. \ bounded(connected_component_set (- S) x)}" by (auto simp: outside_def) (metis Compl_iff bounded_empty connected_component_eq_empty) lemma inside_no_overlap [simp]: "inside S \ S = {}" @@ -2619,7 +2619,7 @@ lemma unbounded_outside: fixes S :: "'a::{real_normed_vector, perfect_space} set" - shows "bounded S \ ~ bounded(outside S)" + shows "bounded S \ \ bounded(outside S)" using cobounded_imp_unbounded cobounded_outside by blast lemma bounded_inside: @@ -2655,7 +2655,7 @@ lemma not_outside_connected_component_lt: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" and "2 \ DIM('a)" - shows "- (outside S) = {x. \B. \y. B < norm(y) \ ~ (connected_component (- S) x y)}" + shows "- (outside S) = {x. \B. \y. B < norm(y) \ \ connected_component (- S) x y}" proof - obtain B::real where B: "0 < B" and Bno: "\x. x \ S \ norm x \ B" using S [simplified bounded_pos] by auto @@ -2681,24 +2681,24 @@ lemma not_outside_connected_component_le: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" "2 \ DIM('a)" - shows "- (outside S) = {x. \B. \y. B \ norm(y) \ ~ (connected_component (- S) x y)}" + shows "- (outside S) = {x. \B. \y. B \ norm(y) \ \ connected_component (- S) x y}" apply (auto intro: less_imp_le simp: not_outside_connected_component_lt [OF assms]) by (meson gt_ex less_le_trans) lemma inside_connected_component_lt: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" "2 \ DIM('a)" - shows "inside S = {x. (x \ S) \ (\B. \y. B < norm(y) \ ~(connected_component (- S) x y))}" + shows "inside S = {x. (x \ S) \ (\B. \y. B < norm(y) \ \ connected_component (- S) x y)}" by (auto simp: inside_outside not_outside_connected_component_lt [OF assms]) lemma inside_connected_component_le: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" "2 \ DIM('a)" - shows "inside S = {x. (x \ S) \ (\B. \y. B \ norm(y) \ ~(connected_component (- S) x y))}" + shows "inside S = {x. (x \ S) \ (\B. \y. B \ norm(y) \ \ connected_component (- S) x y)}" by (auto simp: inside_outside not_outside_connected_component_le [OF assms]) lemma inside_subset: - assumes "connected U" and "~bounded U" and "T \ U = - S" + assumes "connected U" and "\ bounded U" and "T \ U = - S" shows "inside S \ T" apply (auto simp: inside_def) by (metis bounded_subset [of "connected_component_set (- S) _"] connected_component_maximal @@ -2794,7 +2794,7 @@ lemma frontier_minimal_separating_closed: fixes S :: "'a::real_normed_vector set" assumes "closed S" - and nconn: "~ connected(- S)" + and nconn: "\ connected(- S)" and C: "C \ components (- S)" and conn: "\T. \closed T; T \ S\ \ connected(- T)" shows "frontier C = S" @@ -3319,7 +3319,7 @@ lemma bounded_unique_outside: fixes s :: "'a :: euclidean_space set" - shows "\bounded s; DIM('a) \ 2\ \ (c \ components (- s) \ ~bounded c \ c = outside s)" + shows "\bounded s; DIM('a) \ 2\ \ (c \ components (- s) \ \ bounded c \ c = outside s)" apply (rule iffI) apply (metis cobounded_unique_unbounded_components connected_outside double_compl outside_bounded_nonempty outside_in_components unbounded_outside) by (simp add: connected_outside outside_bounded_nonempty outside_in_components unbounded_outside) @@ -3365,7 +3365,7 @@ apply (rule le_no) using w wy oint by (force simp: imageI image_mono interiorI interior_subset frontier_def y) - have **: "(~(b \ (- S) = {}) \ ~(b - (- S) = {}) \ (b \ f \ {})) + have **: "(b \ (- S) \ {} \ b - (- S) \ {} \ b \ f \ {}) \ (b \ S \ {}) \ b \ f = {} \ b \ S" for b f and S :: "'b set" by blast @@ -5065,7 +5065,7 @@ done qed -subsection\Sort of induction principle for connected sets\ +subsection\An induction principle for connected sets\ proposition connected_induction: assumes "connected S" @@ -5084,7 +5084,7 @@ done have 2: "openin (subtopology euclidean S) {b. \T. openin (subtopology euclidean S) T \ - b \ T \ (\x\T. P x \ ~ Q x)}" + b \ T \ (\x\T. P x \ \ Q x)}" apply (subst openin_subopen, clarify) apply (rule_tac x=T in exI, auto) done @@ -5742,7 +5742,7 @@ qed blast -subsection\Important special cases of local connectedness and path connectedness\ +subsection\Special cases of local connectedness and path connectedness\ lemma locally_connected_1: assumes @@ -7155,7 +7155,7 @@ lemma connected_card_eq_iff_nontrivial: fixes S :: "'a::metric_space set" - shows "connected S \ uncountable S \ ~(\a. S \ {a})" + shows "connected S \ uncountable S \ \(\a. S \ {a})" apply (auto simp: countable_finite finite_subset) by (metis connected_uncountable is_singletonI' is_singleton_the_elem subset_singleton_iff) @@ -7184,7 +7184,7 @@ proposition path_connected_convex_diff_countable: fixes U :: "'a::euclidean_space set" - assumes "convex U" "~ collinear U" "countable S" + assumes "convex U" "\ collinear U" "countable S" shows "path_connected(U - S)" proof (clarsimp simp add: path_connected_def) fix a b @@ -7202,7 +7202,7 @@ have "?m \ U" using \a \ U\ \b \ U\ \convex U\ convex_contains_segment by force obtain c where "c \ U" and nc_abc: "\ collinear {a,b,c}" - by (metis False \a \ U\ \b \ U\ \~ collinear U\ collinear_triples insert_absorb) + by (metis False \a \ U\ \b \ U\ \\ collinear U\ collinear_triples insert_absorb) have ncoll_mca: "\ collinear {?m,c,a}" by (metis (full_types) \a \ ?m\ collinear_3_trans collinear_midpoint insert_commute nc_abc) have ncoll_mcb: "\ collinear {?m,c,b}" @@ -7289,7 +7289,7 @@ corollary connected_convex_diff_countable: fixes U :: "'a::euclidean_space set" - assumes "convex U" "~ collinear U" "countable S" + assumes "convex U" "\ collinear U" "countable S" shows "connected(U - S)" by (simp add: assms path_connected_convex_diff_countable path_connected_imp_connected) @@ -7344,7 +7344,7 @@ proposition path_connected_openin_diff_countable: fixes S :: "'a::euclidean_space set" assumes "connected S" and ope: "openin (subtopology euclidean (affine hull S)) S" - and "~ collinear S" "countable T" + and "\ collinear S" "countable T" shows "path_connected(S - T)" proof (clarsimp simp add: path_connected_component) fix x y @@ -7360,8 +7360,8 @@ by (auto simp: openin_contains_ball) with \x \ U\ have x: "x \ ball x r \ affine hull S" by auto - have "~ S \ {x}" - using \~ collinear S\ collinear_subset by blast + have "\ S \ {x}" + using \\ collinear S\ collinear_subset by blast then obtain x' where "x' \ x" "x' \ S" by blast obtain y where y: "y \ x" "y \ ball x r \ affine hull S" @@ -7414,7 +7414,7 @@ corollary connected_openin_diff_countable: fixes S :: "'a::euclidean_space set" assumes "connected S" and ope: "openin (subtopology euclidean (affine hull S)) S" - and "~ collinear S" "countable T" + and "\ collinear S" "countable T" shows "connected(S - T)" by (metis path_connected_imp_connected path_connected_openin_diff_countable [OF assms]) @@ -7445,7 +7445,7 @@ -subsection\Self-homeomorphisms shuffling points about in various ways\ +subsection\Self-homeomorphisms shuffling points about\ subsubsection%unimportant\The theorem \homeomorphism_moving_points_exists\\ @@ -7610,7 +7610,7 @@ assumes "affine T" "a \ T" and ST: "ball a r \ T \ S" "S \ T" and u: "u \ ball a r \ T" and v: "v \ ball a r \ T" obtains f g where "homeomorphism S S f g" - "f u = v" "{x. ~ (f x = x \ g x = x)} \ ball a r \ T" + "f u = v" "{x. \ (f x = x \ g x = x)} \ ball a r \ T" proof - obtain f g where hom: "homeomorphism (cball a r \ T) (cball a r \ T) f g" and "f u = v" and fid: "\x. \x \ sphere a r; x \ T\ \ f x = x" @@ -7690,14 +7690,14 @@ and TS: "T \ affine hull S" and S: "connected S" "a \ S" "b \ S" obtains f g where "homeomorphism T T f g" "f a = b" - "{x. ~ (f x = x \ g x = x)} \ S" - "bounded {x. ~ (f x = x \ g x = x)}" + "{x. \ (f x = x \ g x = x)} \ S" + "bounded {x. \ (f x = x \ g x = x)}" proof - have 1: "\h k. homeomorphism T T h k \ h (f d) = d \ - {x. ~ (h x = x \ k x = x)} \ S \ bounded {x. ~ (h x = x \ k x = x)}" + {x. \ (h x = x \ k x = x)} \ S \ bounded {x. \ (h x = x \ k x = x)}" if "d \ S" "f d \ S" and homfg: "homeomorphism T T f g" - and S: "{x. ~ (f x = x \ g x = x)} \ S" - and bo: "bounded {x. ~ (f x = x \ g x = x)}" for d f g + and S: "{x. \ (f x = x \ g x = x)} \ S" + and bo: "bounded {x. \ (f x = x \ g x = x)}" for d f g proof (intro exI conjI) show homgf: "homeomorphism T T g f" by (metis homeomorphism_symD homfg) @@ -7722,7 +7722,7 @@ by force show "{x. \ ((f2 \ f1) x = x \ (g1 \ g2) x = x)} \ S" using sub by force - have "bounded ({x. ~(f1 x = x \ g1 x = x)} \ {x. ~(f2 x = x \ g2 x = x)})" + have "bounded ({x. \(f1 x = x \ g1 x = x)} \ {x. \(f2 x = x \ g2 x = x)})" using bo by simp then show "bounded {x. \ ((f2 \ f1) x = x \ (g1 \ g2) x = x)}" by (rule bounded_subset) auto @@ -7753,7 +7753,7 @@ done qed have "\f g. homeomorphism T T f g \ f a = b \ - {x. ~ (f x = x \ g x = x)} \ S \ bounded {x. ~ (f x = x \ g x = x)}" + {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" apply (rule connected_equivalence_relation [OF S], safe) apply (blast intro: 1 2 3)+ done @@ -7769,7 +7769,7 @@ and ope: "openin (subtopology euclidean (affine hull S)) S" and "S \ T" "T \ affine hull S" "connected S" shows "\f g. homeomorphism T T f g \ (\i \ K. f(x i) = y i) \ - {x. ~ (f x = x \ g x = x)} \ S \ bounded {x. ~ (f x = x \ g x = x)}" + {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" using assms proof (induction K) case empty @@ -7783,11 +7783,11 @@ and xyS: "\i. i \ K \ x i \ S \ y i \ S" by (simp_all add: pairwise_insert) obtain f g where homfg: "homeomorphism T T f g" and feq: "\i. i \ K \ f(x i) = y i" - and fg_sub: "{x. ~ (f x = x \ g x = x)} \ S" - and bo_fg: "bounded {x. ~ (f x = x \ g x = x)}" + and fg_sub: "{x. \ (f x = x \ g x = x)} \ S" + and bo_fg: "bounded {x. \ (f x = x \ g x = x)}" using insert.IH [OF xyS pw] insert.prems by (blast intro: that) then have "\f g. homeomorphism T T f g \ (\i \ K. f(x i) = y i) \ - {x. ~ (f x = x \ g x = x)} \ S \ bounded {x. ~ (f x = x \ g x = x)}" + {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" using insert by blast have aff_eq: "affine hull (S - y ` K) = affine hull S" apply (rule affine_hull_Diff) @@ -7832,7 +7832,7 @@ using feq hk_sub by (auto simp: heq) show "{x. \ ((h \ f) x = x \ (g \ k) x = x)} \ S" using fg_sub hk_sub by force - have "bounded ({x. ~(f x = x \ g x = x)} \ {x. ~(h x = x \ k x = x)})" + have "bounded ({x. \(f x = x \ g x = x)} \ {x. \(h x = x \ k x = x)})" using bo_fg bo_hk bounded_Un by blast then show "bounded {x. \ ((h \ f) x = x \ (g \ k) x = x)}" by (rule bounded_subset) auto @@ -7846,7 +7846,7 @@ and pw: "pairwise (\i j. (x i \ x j) \ (y i \ y j)) K" and S: "S \ T" "T \ affine hull S" "connected S" obtains f g where "homeomorphism T T f g" "\i. i \ K \ f(x i) = y i" - "{x. ~ (f x = x \ g x = x)} \ S" "bounded {x. (~ (f x = x \ g x = x))}" + "{x. \ (f x = x \ g x = x)} \ S" "bounded {x. (\ (f x = x \ g x = x))}" proof (cases "S = {}") case True then show ?thesis @@ -7995,8 +7995,8 @@ fixes T :: "real set" assumes "open U" "open S" "connected S" "U \ {}" "finite K" "K \ S" "U \ S" "S \ T" obtains f g where "homeomorphism T T f g" - "\x. x \ K \ f x \ U" "{x. (~ (f x = x \ g x = x))} \ S" - "bounded {x. (~ (f x = x \ g x = x))}" + "\x. x \ K \ f x \ U" "{x. (\ (f x = x \ g x = x))} \ S" + "bounded {x. (\ (f x = x \ g x = x))}" proof - obtain c d where "box c d \ {}" "cbox c d \ U" proof - @@ -8109,8 +8109,8 @@ proposition homeomorphism_grouping_points_exists: fixes S :: "'a::euclidean_space set" assumes "open U" "open S" "connected S" "U \ {}" "finite K" "K \ S" "U \ S" "S \ T" - obtains f g where "homeomorphism T T f g" "{x. (~ (f x = x \ g x = x))} \ S" - "bounded {x. (~ (f x = x \ g x = x))}" "\x. x \ K \ f x \ U" + obtains f g where "homeomorphism T T f g" "{x. (\ (f x = x \ g x = x))} \ S" + "bounded {x. (\ (f x = x \ g x = x))}" "\x. x \ K \ f x \ U" proof (cases "2 \ DIM('a)") case True have TS: "T \ affine hull S" @@ -8172,10 +8172,10 @@ using sub hj apply (drule_tac c="h x" in subsetD, force) by (metis imageE) - have "bounded (j ` {x. (~ (f x = x \ g x = x))})" + have "bounded (j ` {x. (\ (f x = x \ g x = x))})" by (rule bounded_linear_image [OF bou]) (use \linear j\ linear_conv_bounded_linear in auto) moreover - have *: "{x. ~((j \ f \ h) x = x \ (j \ g \ h) x = x)} = j ` {x. (~ (f x = x \ g x = x))}" + have *: "{x. \((j \ f \ h) x = x \ (j \ g \ h) x = x)} = j ` {x. (\ (f x = x \ g x = x))}" using hj by (auto simp: jf jg image_iff, metis+) ultimately show "bounded {x. \ ((j \ f \ h) x = x \ (j \ g \ h) x = x)}" by metis @@ -8190,8 +8190,8 @@ assumes opeU: "openin (subtopology euclidean S) U" and opeS: "openin (subtopology euclidean (affine hull S)) S" and "U \ {}" "finite K" "K \ S" and S: "S \ T" "T \ affine hull S" "connected S" - obtains f g where "homeomorphism T T f g" "{x. (~ (f x = x \ g x = x))} \ S" - "bounded {x. (~ (f x = x \ g x = x))}" "\x. x \ K \ f x \ U" + obtains f g where "homeomorphism T T f g" "{x. (\ (f x = x \ g x = x))} \ S" + "bounded {x. (\ (f x = x \ g x = x))}" "\x. x \ K \ f x \ U" proof (cases "2 \ aff_dim S") case True have opeU': "openin (subtopology euclidean (affine hull S)) U" @@ -8208,7 +8208,7 @@ then obtain \ where \: "bij_betw \ K P" using \finite K\ finite_same_card_bij by blast have "\f g. homeomorphism T T f g \ (\i \ K. f(id i) = \ i) \ - {x. ~ (f x = x \ g x = x)} \ S \ bounded {x. ~ (f x = x \ g x = x)}" + {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" proof (rule homeomorphism_moving_points_exists_gen [OF \finite K\ _ _ True opeS S]) show "\i. i \ K \ id i \ S \ \ i \ S" by (metis id_apply opeU openin_contains_cball subsetCE \P \ U\ \bij_betw \ K P\ \K \ S\ bij_betwE) @@ -8314,10 +8314,10 @@ next have "compact (j ` closure {x. \ (f x = x \ g x = x)})" using bou by (auto simp: compact_continuous_image cont_hj) - then have "bounded (j ` {x. (~ (f x = x \ g x = x))})" + then have "bounded (j ` {x. \ (f x = x \ g x = x)})" by (rule bounded_closure_image [OF compact_imp_bounded]) moreover - have *: "{x \ affine hull S. j (f (h x)) \ x \ j (g (h x)) \ x} = j ` {x. (~ (f x = x \ g x = x))}" + have *: "{x \ affine hull S. j (f (h x)) \ x \ j (g (h x)) \ x} = j ` {x. (\ (f x = x \ g x = x))}" using h j by (auto simp: image_iff; metis) ultimately have "bounded {x \ affine hull S. j (f (h x)) \ x \ j (g (h x)) \ x}" by metis @@ -8337,7 +8337,7 @@ qed qed -subsection\nullhomotopic mappings\ +subsection\Nullhomotopic mappings\ text\ A mapping out of a sphere is nullhomotopic iff it extends to the ball. This even works out in the degenerate cases when the radius is \\\ 0, and diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Poly_Roots.thy --- a/src/HOL/Analysis/Poly_Roots.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Poly_Roots.thy Thu Dec 27 19:48:41 2018 +0100 @@ -169,7 +169,7 @@ by auto have c0: "c 0 = - (a * b 0)" using b [of 0] by simp - then have extr_prem: "~ (\k\n. b k \ 0) \ \k. k \ 0 \ k \ Suc n \ c k \ 0" + 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) @@ -205,7 +205,7 @@ shows "(\z. (\i\n. c i * z^i) = 0) \ (\k. k \ n \ c k = 0)" proof (cases "(\z. (\i\n. c i * z^i) = 0)") case True - then have "~ finite {z. (\i\n. c i * z^i) = 0}" + then have "\ finite {z. (\i\n. c i * z^i) = 0}" by (simp add: infinite_UNIV_char_0) with True show ?thesis by (metis (poly_guards_query) polyfun_rootbound_finite) diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Polytope.thy --- a/src/HOL/Analysis/Polytope.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Polytope.thy Thu Dec 27 19:48:41 2018 +0100 @@ -270,7 +270,7 @@ lemma%unimportant subset_of_face_of_affine_hull: fixes S :: "'a::euclidean_space set" - assumes T: "T face_of S" and "convex S" "U \ S" and dis: "~disjnt (affine hull T) (rel_interior U)" + assumes T: "T face_of S" and "convex S" "U \ S" and dis: "\ disjnt (affine hull T) (rel_interior U)" shows "U \ T" apply (rule subset_of_face_of [OF T \U \ S\]) using face_of_imp_eq_affine_Int [OF \convex S\ T] @@ -846,10 +846,10 @@ using affine_parallel_slice affine_affine_hull by metis show ?thesis proof (intro conjI impI allI ballI exI) - have *: "S \ - (affine hull S \ {x. P x}) \ affine hull S \ {x. Q x} \ S \ {x. ~P x \ Q x}" + have *: "S \ - (affine hull S \ {x. P x}) \ affine hull S \ {x. Q x} \ S \ {x. \ P x \ Q x}" for P Q using hull_subset by fastforce - have "S \ {x. ~ (a' \ x \ b') \ a' \ x = b'}" + have "S \ {x. \ (a' \ x \ b') \ a' \ x = b'}" apply (rule *) apply (simp only: le eq) using Ssub by auto @@ -919,7 +919,7 @@ "{x. x extreme_point_of (convex hull S)} \ S" using%unimportant extreme_point_of_convex_hull by auto -lemma%unimportant extreme_point_of_empty [simp]: "~ (x extreme_point_of {})" +lemma%unimportant extreme_point_of_empty [simp]: "\ (x extreme_point_of {})" by (simp add: extreme_point_of_def) lemma%unimportant extreme_point_of_singleton [iff]: "x extreme_point_of {a} \ x = a" @@ -974,10 +974,10 @@ (infixr "(facet'_of)" 50) where "F facet_of S \ F face_of S \ F \ {} \ aff_dim F = aff_dim S - 1" -lemma%unimportant facet_of_empty [simp]: "~ S facet_of {}" +lemma%unimportant facet_of_empty [simp]: "\ S facet_of {}" by (simp add: facet_of_def) -lemma%unimportant facet_of_irrefl [simp]: "~ S facet_of S " +lemma%unimportant facet_of_irrefl [simp]: "\ S facet_of S " by (simp add: facet_of_def) lemma%unimportant facet_of_imp_face_of: "F facet_of S \ F face_of S" @@ -1088,7 +1088,7 @@ using * [of "norm b" "norm a" "1-u" for u] noax that apply (simp add: add.commute) done - ultimately have "~ (norm a < norm x) \ ~ (norm b < norm x)" + ultimately have "\ (norm a < norm x) \ \ (norm b < norm x)" by auto then show ?thesis using different_norm_3_collinear_points noax nobx that(3) by fastforce @@ -1268,7 +1268,7 @@ lemma%unimportant extreme_point_of_convex_hull_affine_independent: fixes S :: "'a::euclidean_space set" shows - "~ affine_dependent S + "\ affine_dependent S \ (x extreme_point_of (convex hull S) \ x \ S)" by (metis aff_independent_finite affine_dependent_def affine_hull_convex_hull extreme_point_of_convex_hull_convex_independent finite_imp_compact hull_inc) @@ -1473,7 +1473,7 @@ proposition%important face_of_convex_hull_affine_independent: fixes S :: "'a::euclidean_space set" - assumes "~ affine_dependent S" + assumes "\ affine_dependent S" shows "(T face_of (convex hull S) \ (\c. c \ S \ T = convex hull c))" (is "?lhs = ?rhs") proof%unimportant @@ -1495,7 +1495,7 @@ lemma%unimportant facet_of_convex_hull_affine_independent: fixes S :: "'a::euclidean_space set" - assumes "~ affine_dependent S" + assumes "\ affine_dependent S" shows "T facet_of (convex hull S) \ T \ {} \ (\u. u \ S \ T = convex hull (S - {u}))" (is "?lhs = ?rhs") @@ -1508,7 +1508,7 @@ by (auto simp: face_of_convex_hull_affine_independent [OF assms]) then have affs: "aff_dim S = aff_dim c + 1" by (metis aff_dim_convex_hull afft eq_diff_eq) - have "~ affine_dependent c" + have "\ affine_dependent c" using \c \ S\ affine_dependent_subset assms by blast with affs have "card (S - c) = 1" apply (simp add: aff_dim_affine_independent [symmetric] aff_dim_convex_hull) @@ -1543,7 +1543,7 @@ lemma%unimportant facet_of_convex_hull_affine_independent_alt: fixes S :: "'a::euclidean_space set" shows - "~affine_dependent S + "\affine_dependent S \ (T facet_of (convex hull S) \ 2 \ card S \ (\u. u \ S \ T = convex hull (S - {u})))" apply (simp add: facet_of_convex_hull_affine_independent) @@ -1897,20 +1897,20 @@ have "\a' b'. a' \ 0 \ affine hull S \ {x. a' \ x \ b'} = affine hull S \ h \ (\w \ affine hull S. (w + a') \ affine hull S)" - if "h \ F" "~(affine hull S \ h)" for h + if "h \ F" "\(affine hull S \ h)" for h proof - have "a h \ 0" and "h = {x. a h \ x \ b h}" "h \ \F = \F" using \h \ F\ ab by auto then have "(affine hull S) \ {x. a h \ x \ b h} \ {}" by (metis (no_types) affine_hull_eq_empty inf.absorb_iff2 inf_assoc inf_bot_right inf_commute seq that(2)) - moreover have "~ (affine hull S \ {x. a h \ x \ b h})" + moreover have "\ (affine hull S \ {x. a h \ x \ b h})" using \h = {x. a h \ x \ b h}\ that(2) by blast ultimately show ?thesis using affine_parallel_slice [of "affine hull S"] by (metis \h = {x. a h \ x \ b h}\ affine_affine_hull) qed then obtain a b - where ab: "\h. \h \ F; ~ (affine hull S \ h)\ + where ab: "\h. \h \ F; \ (affine hull S \ h)\ \ a h \ 0 \ affine hull S \ {x. a h \ x \ b h} = affine hull S \ h \ (\w \ affine hull S. (w + a h) \ affine hull S)" @@ -1918,7 +1918,7 @@ have seq2: "S = affine hull S \ (\h\{h \ F. \ affine hull S \ h}. {x. a h \ x \ b h})" by (subst seq) (auto simp: ab INT_extend_simps) show ?thesis - apply (rule_tac x="(\h. {x. a h \ x \ b h}) ` {h. h \ F \ ~(affine hull S \ h)}" in exI) + apply (rule_tac x="(\h. {x. a h \ x \ b h}) ` {h. h \ F \ \(affine hull S \ h)}" in exI) apply (intro conjI seq2) using \finite F\ apply force using ab apply blast @@ -1958,9 +1958,9 @@ done then obtain F where F: "card F = n" "finite F" and seq: "?P F" and aff: "?Q F" by blast - then have "~ (finite g \ ?P g \ ?Q g)" if "card g < n" for g + then have "\ (finite g \ ?P g \ ?Q g)" if "card g < n" for g using that by (auto simp: n_def dest!: not_less_Least) - then have *: "~ (?P g \ ?Q g)" if "g \ F" for g + then have *: "\ (?P g \ ?Q g)" if "g \ F" for g using that \finite F\ psubset_card_mono \card F = n\ by (metis finite_Int inf.strict_order_iff) have 1: "\F'. F' \ F \ S \ affine hull S \ \F'" @@ -2230,7 +2230,7 @@ using seq by blast have "F \ {}" using assms by (metis affine_Int affine_Inter affine_affine_hull ex_in_conv face_of_affine_trivial) - then obtain i where "i \ F" "~ (a i \ x < b i)" + then obtain i where "i \ F" "\ (a i \ x < b i)" using \x \ S\ rels xnot by auto with xint have "a i \ x = b i" by (metis eq_iff mem_Collect_eq not_le Inter_iff faceq) @@ -2745,7 +2745,7 @@ lemma rel_boundary_of_convex_hull: fixes S :: "'a::euclidean_space set" - assumes "~ affine_dependent S" + assumes "\ affine_dependent S" shows "(convex hull S) - rel_interior(convex hull S) = (\a\S. convex hull (S - {a}))" proof - have "finite S" by (metis assms aff_independent_finite) @@ -3074,17 +3074,17 @@ text\The notion of n-simplex for integer @{term"n \ -1"}\ definition simplex :: "int \ 'a::euclidean_space set \ bool" (infix "simplex" 50) - where "n simplex S \ \C. ~(affine_dependent C) \ int(card C) = n + 1 \ S = convex hull C" + where "n simplex S \ \C. \ affine_dependent C \ int(card C) = n + 1 \ S = convex hull C" lemma simplex: "n simplex S \ (\C. finite C \ - ~(affine_dependent C) \ + \ affine_dependent C \ int(card C) = n + 1 \ S = convex hull C)" by (auto simp add: simplex_def intro: aff_independent_finite) lemma simplex_convex_hull: - "~affine_dependent C \ int(card C) = n + 1 \ n simplex (convex hull C)" + "\ affine_dependent C \ int(card C) = n + 1 \ n simplex (convex hull C)" by (auto simp add: simplex_def) lemma convex_simplex: "n simplex S \ convex S" @@ -3170,7 +3170,7 @@ assumes "n simplex S" and a: "a \ affine hull S" shows "(n+1) simplex (convex hull (insert a S))" proof - - obtain C where C: "finite C" "~(affine_dependent C)" "int(card C) = n+1" and S: "S = convex hull C" + obtain C where C: "finite C" "\ affine_dependent C" "int(card C) = n+1" and S: "S = convex hull C" using assms unfolding simplex by force show ?thesis unfolding simplex @@ -3270,11 +3270,11 @@ U [unfolded rel_frontier_def] tnot by (auto simp: closed_segment_eq_open) ultimately - have "~(between (t,u) z | between (u,z) t | between (z,t) u)" if "x \ z" + have "\(between (t,u) z | between (u,z) t | between (z,t) u)" if "x \ z" using that xt xu apply (simp add: between_mem_segment [symmetric]) by (metis between_commute between_trans_2 between_antisym) - then have "~ collinear {t, z, u}" if "x \ z" + then have "\ collinear {t, z, u}" if "x \ z" by (auto simp: that collinear_between_cases between_commute) moreover have "collinear {t, z, x}" by (metis closed_segment_commute collinear_2 collinear_closed_segment collinear_triples ends_in_segment(1) insert_absorb insert_absorb2 xt) @@ -3352,7 +3352,7 @@ by (auto simp: \_def poly\ polytope_imp_convex polytope_imp_closed) have in_rel_interior: "(SOME z. z \ rel_interior C) \ rel_interior C" if "C \ \" for C using that poly\ polytope_imp_convex rel_interior_aff_dim some_in_eq by (fastforce simp: \_def) - have *: "\T. ~affine_dependent T \ card T \ n \ aff_dim K < n \ K = convex hull T" + have *: "\T. \ affine_dependent T \ card T \ n \ aff_dim K < n \ K = convex hull T" if "K \ \" for K proof - obtain r where r: "r simplex K" diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Riemann_Mapping.thy --- a/src/HOL/Analysis/Riemann_Mapping.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Riemann_Mapping.thy Thu Dec 27 19:48:41 2018 +0100 @@ -860,7 +860,7 @@ lemma frontier_properties: assumes "simply_connected S" shows "if bounded S then connected(frontier S) - else \C \ components(frontier S). ~bounded C" + else \C \ components(frontier S). \ bounded C" proof - have "S = {} \ S homeomorphic ball (0::complex) 1" using simply_connected_eq_homeomorphic_to_disc assms openS by blast @@ -1217,7 +1217,7 @@ qed lemma empty_inside: - assumes "connected S" "\C. C \ components (- S) \ ~bounded C" + assumes "connected S" "\C. C \ components (- S) \ \ bounded C" shows "inside S = {}" using assms by (auto simp: components_def inside_def) diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Starlike.thy --- a/src/HOL/Analysis/Starlike.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Starlike.thy Thu Dec 27 19:48:41 2018 +0100 @@ -3726,7 +3726,7 @@ lemma affine_independent_convex_affine_hull: fixes s :: "'a::euclidean_space set" - assumes "~affine_dependent s" "t \ s" + assumes "\ affine_dependent s" "t \ s" shows "convex hull t = affine hull t \ convex hull s" proof - have fin: "finite s" "finite t" using assms aff_independent_finite finite_subset by auto @@ -3764,7 +3764,7 @@ lemma affine_independent_span_eq: fixes s :: "'a::euclidean_space set" - assumes "~affine_dependent s" "card s = Suc (DIM ('a))" + assumes "\ affine_dependent s" "card s = Suc (DIM ('a))" shows "affine hull s = UNIV" proof (cases "s = {}") case True then show ?thesis @@ -3789,7 +3789,7 @@ lemma affine_independent_span_gt: fixes s :: "'a::euclidean_space set" - assumes ind: "~ affine_dependent s" and dim: "DIM ('a) < card s" + assumes ind: "\ affine_dependent s" and dim: "DIM ('a) < card s" shows "affine hull s = UNIV" apply (rule affine_independent_span_eq [OF ind]) apply (rule antisym) @@ -3831,7 +3831,7 @@ lemma rel_interior_convex_hull_explicit: fixes s :: "'a::euclidean_space set" - assumes "~ affine_dependent s" + assumes "\ affine_dependent s" shows "rel_interior(convex hull s) = {y. \u. (\x \ s. 0 < u x) \ sum u s = 1 \ sum (\x. u x *\<^sub>R x) s = y}" (is "?lhs = ?rhs") @@ -3904,18 +3904,18 @@ lemma interior_convex_hull_explicit_minimal: fixes s :: "'a::euclidean_space set" shows - "~ affine_dependent s + "\ affine_dependent s ==> interior(convex hull s) = (if card(s) \ DIM('a) then {} else {y. \u. (\x \ s. 0 < u x) \ sum u s = 1 \ (\x\s. u x *\<^sub>R x) = y})" apply (simp add: aff_independent_finite empty_interior_convex_hull, clarify) apply (rule trans [of _ "rel_interior(convex hull s)"]) - apply (simp add: affine_hull_convex_hull affine_independent_span_gt rel_interior_interior) + apply (simp add: affine_independent_span_gt rel_interior_interior) by (simp add: rel_interior_convex_hull_explicit) lemma interior_convex_hull_explicit: fixes s :: "'a::euclidean_space set" - assumes "~ affine_dependent s" + assumes "\ affine_dependent s" shows "interior(convex hull s) = (if card(s) \ DIM('a) then {} @@ -4082,7 +4082,7 @@ lemma rel_frontier_convex_hull_explicit: fixes s :: "'a::euclidean_space set" - assumes "~ affine_dependent s" + assumes "\ affine_dependent s" shows "rel_frontier(convex hull s) = {y. \u. (\x \ s. 0 \ u x) \ (\x \ s. u x = 0) \ sum u s = 1 \ sum (\x. u x *\<^sub>R x) s = y}" proof - @@ -4106,7 +4106,7 @@ lemma frontier_convex_hull_explicit: fixes s :: "'a::euclidean_space set" - assumes "~ affine_dependent s" + assumes "\ affine_dependent s" shows "frontier(convex hull s) = {y. \u. (\x \ s. 0 \ u x) \ (DIM ('a) < card s \ (\x \ s. u x = 0)) \ sum u s = 1 \ sum (\x. u x *\<^sub>R x) s = y}" @@ -4133,7 +4133,7 @@ lemma rel_frontier_convex_hull_cases: fixes s :: "'a::euclidean_space set" - assumes "~ affine_dependent s" + assumes "\ affine_dependent s" shows "rel_frontier(convex hull s) = \{convex hull (s - {x}) |x. x \ s}" proof - have fs: "finite s" @@ -4164,7 +4164,7 @@ lemma frontier_convex_hull_eq_rel_frontier: fixes s :: "'a::euclidean_space set" - assumes "~ affine_dependent s" + assumes "\ affine_dependent s" shows "frontier(convex hull s) = (if card s \ DIM ('a) then convex hull s else rel_frontier(convex hull s))" using assms @@ -4174,7 +4174,7 @@ lemma frontier_convex_hull_cases: fixes s :: "'a::euclidean_space set" - assumes "~ affine_dependent s" + assumes "\ affine_dependent s" shows "frontier(convex hull s) = (if card s \ DIM ('a) then convex hull s else \{convex hull (s - {x}) |x. x \ s})" by (simp add: assms frontier_convex_hull_eq_rel_frontier rel_frontier_convex_hull_cases) @@ -4668,7 +4668,7 @@ lemma interior_convex_hull_3_minimal: fixes a :: "'a::euclidean_space" - shows "\~ collinear{a,b,c}; DIM('a) = 2\ + shows "\\ collinear{a,b,c}; DIM('a) = 2\ \ interior(convex hull {a,b,c}) = {v. \x y z. 0 < x \ 0 < y \ 0 < z \ x + y + z = 1 \ x *\<^sub>R a + y *\<^sub>R b + z *\<^sub>R c = v}" @@ -5862,7 +5862,7 @@ lemma affine_dependent_choose: fixes a :: "'a :: euclidean_space" - assumes "~(affine_dependent S)" + assumes "\(affine_dependent S)" shows "affine_dependent(insert a S) \ a \ S \ a \ affine hull S" (is "?lhs = ?rhs") proof safe @@ -5887,7 +5887,7 @@ lemma affine_independent_insert: fixes a :: "'a :: euclidean_space" - shows "\~(affine_dependent S); a \ affine hull S\ \ ~(affine_dependent(insert a S))" + shows "\\ affine_dependent S; a \ affine hull S\ \ \ affine_dependent(insert a S)" by (simp add: affine_dependent_choose) lemma subspace_bounded_eq_trivial: @@ -6158,7 +6158,7 @@ lemma fixes s :: "'a::euclidean_space set" - assumes "~ (affine_dependent(s \ t))" + assumes "\ affine_dependent(s \ t)" shows convex_hull_Int_subset: "convex hull s \ convex hull t \ convex hull (s \ t)" (is ?C) and affine_hull_Int_subset: "affine hull s \ affine hull t \ affine hull (s \ t)" (is ?A) proof - @@ -6201,7 +6201,7 @@ proposition affine_hull_Int: fixes s :: "'a::euclidean_space set" - assumes "~ (affine_dependent(s \ t))" + assumes "\ affine_dependent(s \ t)" shows "affine hull (s \ t) = affine hull s \ affine hull t" apply (rule subset_antisym) apply (simp add: hull_mono) @@ -6209,7 +6209,7 @@ proposition convex_hull_Int: fixes s :: "'a::euclidean_space set" - assumes "~ (affine_dependent(s \ t))" + assumes "\ affine_dependent(s \ t)" shows "convex hull (s \ t) = convex hull s \ convex hull t" apply (rule subset_antisym) apply (simp add: hull_mono) @@ -6217,7 +6217,7 @@ proposition fixes s :: "'a::euclidean_space set set" - assumes "~ (affine_dependent (\s))" + assumes "\ affine_dependent (\s)" shows affine_hull_Inter: "affine hull (\s) = (\t\s. affine hull t)" (is "?A") and convex_hull_Inter: "convex hull (\s) = (\t\s. convex hull t)" (is "?C") proof - @@ -6247,7 +6247,7 @@ proposition in_convex_hull_exchange_unique: fixes S :: "'a::euclidean_space set" - assumes naff: "~ affine_dependent S" and a: "a \ convex hull S" + assumes naff: "\ affine_dependent S" and a: "a \ convex hull S" and S: "T \ S" "T' \ S" and x: "x \ convex hull (insert a T)" and x': "x \ convex hull (insert a T')" @@ -6395,7 +6395,7 @@ corollary convex_hull_exchange_Int: fixes a :: "'a::euclidean_space" - assumes "~ affine_dependent S" "a \ convex hull S" "T \ S" "T' \ S" + assumes "\ affine_dependent S" "a \ convex hull S" "T \ S" "T' \ S" shows "(convex hull (insert a T)) \ (convex hull (insert a T')) = convex hull (insert a (T \ T'))" apply (rule subset_antisym) @@ -6404,7 +6404,7 @@ lemma Int_closed_segment: fixes b :: "'a::euclidean_space" - assumes "b \ closed_segment a c \ ~collinear{a,b,c}" + assumes "b \ closed_segment a c \ \ collinear{a,b,c}" shows "closed_segment a b \ closed_segment b c = {b}" proof (cases "c = a") case True @@ -7283,7 +7283,7 @@ fixes S :: "'a :: euclidean_space set" assumes "affine S" and "S \ {x. a \ x \ b} \ {}" - and "~ (S \ {x. a \ x \ b})" + and "\ (S \ {x. a \ x \ b})" obtains a' b' where "a' \ 0" "S \ {x. a' \ x \ b'} = S \ {x. a \ x \ b}" "S \ {x. a' \ x = b'} = S \ {x. a \ x = b}" @@ -7669,7 +7669,7 @@ assumes opS: "openin (subtopology euclidean (S \ T)) S" and opT: "openin (subtopology euclidean (S \ T)) T" and contf: "continuous_on S f" and contg: "continuous_on T g" - and fg: "\x. x \ S \ ~P x \ x \ T \ P x \ f x = g x" + and fg: "\x. x \ S \ \P x \ x \ T \ P x \ f x = g x" shows "continuous_on (S \ T) (\x. if P x then f x else g x)" proof - have "\x. x \ S \ (if P x then f x else g x) = f x" "\x. x \ T \ (if P x then f x else g x) = g x" diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Tagged_Division.thy --- a/src/HOL/Analysis/Tagged_Division.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Tagged_Division.thy Thu Dec 27 19:48:41 2018 +0100 @@ -1360,7 +1360,7 @@ and k: "k \ Basis" shows "division_points (cbox a b \ {x. x\k \ c}) {l \ {x. x\k \ c} | l . l \ d \ l \ {x. x\k \ c} \ {}} \ division_points (cbox a b) d" (is ?t1) - and "division_points (cbox a b \ {x. x\k \ c}) {l \ {x. x\k \ c} | l . l \ d \ ~(l \ {x. x\k \ c} = {})} \ + and "division_points (cbox a b \ {x. x\k \ c}) {l \ {x. x\k \ c} | l . l \ d \ \(l \ {x. x\k \ c} = {})} \ division_points (cbox a b) d" (is ?t2) proof%unimportant - note assm = division_ofD[OF assms(1)] @@ -2536,7 +2536,7 @@ show "\K. K \ ?D1 \ finite {L. L \ ?D1 \ K \ L}" using fin by simp (metis (mono_tags, lifting) Collect_mono rev_finite_subset) qed - let ?\ = "{K \ \. \K'. K' \ \ \ K \ K' \ ~(K \ K')}" + let ?\ = "{K \ \. \K'. K' \ \ \ K \ K' \ \(K \ K')}" show ?thesis proof (rule that) show "countable ?\" diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Analysis/Topology_Euclidean_Space.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy Thu Dec 27 19:48:41 2018 +0100 @@ -1104,7 +1104,7 @@ lemma connected_openin: "connected S \ - ~(\E1 E2. openin (subtopology euclidean S) E1 \ + \(\E1 E2. openin (subtopology euclidean S) E1 \ openin (subtopology euclidean S) E2 \ S \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" apply (simp add: connected_def openin_open disjoint_iff_not_equal, safe) @@ -1113,7 +1113,7 @@ lemma connected_openin_eq: "connected S \ - ~(\E1 E2. openin (subtopology euclidean S) E1 \ + \(\E1 E2. openin (subtopology euclidean S) E1 \ openin (subtopology euclidean S) E2 \ E1 \ E2 = S \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" @@ -1152,7 +1152,7 @@ lemma connected_closedin_eq: "connected S \ - ~(\E1 E2. + \(\E1 E2. closedin (subtopology euclidean S) E1 \ closedin (subtopology euclidean S) E2 \ E1 \ E2 = S \ E1 \ E2 = {} \ @@ -1946,7 +1946,7 @@ by (auto simp: box_def inner_sum_left inner_Basis sum.If_cases) qed -subsection \Intervals in general, including infinite and mixtures of open and closed\ +subsection \General Intervals\ definition%important "is_interval (s::('a::euclidean_space) set) \ (\a\s. \b\s. \x. (\i\Basis. ((a\i \ x\i \ x\i \ b\i) \ (b\i \ x\i \ x\i \ a\i))) \ x \ s)" @@ -3645,7 +3645,7 @@ corollary cobounded_imp_unbounded: fixes S :: "'a::{real_normed_vector, perfect_space} set" - shows "bounded (- S) \ ~ (bounded S)" + shows "bounded (- S) \ \ bounded S" using bounded_Un [of S "-S"] by (simp add: sup_compl_top) lemma bounded_dist_comp: @@ -4743,7 +4743,7 @@ lemma not_compact_UNIV[simp]: fixes s :: "'a::{real_normed_vector,perfect_space,heine_borel} set" - shows "~ compact (UNIV::'a set)" + shows "\ compact (UNIV::'a set)" by (simp add: compact_eq_bounded_closed) text\Representing sets as the union of a chain of compact sets.\ diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Weierstrass_Theorems.thy --- a/src/HOL/Analysis/Weierstrass_Theorems.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Weierstrass_Theorems.thy Thu Dec 27 19:48:41 2018 +0100 @@ -767,7 +767,7 @@ assumes "compact S" "\c. P(\x. c::real)" "\f. P f \ continuous_on S f" "\f g. P(f) \ P(g) \ P(\x. f x + g x)" "\f g. P(f) \ P(g) \ P(\x. f x * g x)" - "\x y. x \ S \ y \ S \ ~(x = y) \ \f. P(f) \ ~(f x = f y)" + "\x y. x \ S \ y \ S \ x \ y \ \f. P(f) \ f x \ f y" "continuous_on S f" "0 < e" shows "\g. P(g) \ (\x \ S. \f x - g x\ < e)" diff -r 04e54f57a869 -r f9bf65d90b69 src/HOL/Analysis/Winding_Numbers.thy --- a/src/HOL/Analysis/Winding_Numbers.thy Thu Dec 27 17:36:19 2018 +0100 +++ b/src/HOL/Analysis/Winding_Numbers.thy Thu Dec 27 19:48:41 2018 +0100 @@ -22,7 +22,7 @@ lemma wn_triangle1: assumes "0 \ interior(convex hull {a,b,c})" - shows "~ (Im(a/b) \ 0 \ 0 \ Im(b/c))" + shows "\ (Im(a/b) \ 0 \ 0 \ Im(b/c))" proof - { assume 0: "Im(a/b) \ 0" "0 \ Im(b/c)" have "0 \ interior (convex hull {a,b,c})" @@ -224,7 +224,7 @@ by (simp_all add: assms path_image_join) with \0 < e\ dac have "c \ affine hull {a - of_real e, a + of_real e}" by (simp add: segment_as_ball not_le) - with \0 < e\ have *: "~collinear{a - e, c,a + e}" + with \0 < e\ have *: "\ collinear {a - e, c,a + e}" using collinear_3_affine_hull [of "a-e" "a+e"] by (auto simp: insert_commute) have 13: "1/3 + 1/3 + 1/3 = (1::real)" by simp have "(1/3) *\<^sub>R (a - of_real e) + (1/3) *\<^sub>R c + (1/3) *\<^sub>R (a + of_real e) \ interior(convex hull {a - e, c, a + e})"