--- 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 \<Rightarrow> bool" where
"connected_space X \<equiv>
- ~(\<exists>E1 E2. openin X E1 \<and> openin X E2 \<and>
+ \<not>(\<exists>E1 E2. openin X E1 \<and> openin X E2 \<and>
topspace X \<subseteq> E1 \<union> E2 \<and> E1 \<inter> E2 = {} \<and> E1 \<noteq> {} \<and> E2 \<noteq> {})"
definition connectedin :: "'a topology \<Rightarrow> 'a set \<Rightarrow> bool" where
@@ -2093,10 +2093,10 @@
lemma connectedin_closedin:
"connectedin X S \<longleftrightarrow>
S \<subseteq> topspace X \<and>
- ~(\<exists>E1 E2. closedin X E1 \<and> closedin X E2 \<and>
+ \<not>(\<exists>E1 E2. closedin X E1 \<and> closedin X E2 \<and>
S \<subseteq> (E1 \<union> E2) \<and>
(E1 \<inter> E2 \<inter> S = {}) \<and>
- ~(E1 \<inter> S = {}) \<and> ~(E2 \<inter> S = {}))"
+ \<not>(E1 \<inter> S = {}) \<and> \<not>(E2 \<inter> S = {}))"
proof -
have *: "(\<exists>E1:: 'a set. \<exists>E2:: 'a set. (\<exists>T1:: 'a set. P1 T1 \<and> E1 = f1 T1) \<and> (\<exists>T2:: 'a set. P2 T2 \<and> E2 = f2 T2) \<and>
R E1 E2) \<longleftrightarrow> (\<exists>T1 T2. P1 T1 \<and> P2 T2 \<and> 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:
- "\<lbrakk>connectedin X S; connectedin X T; ~separatedin X S T\<rbrakk> \<Longrightarrow> connectedin X (S \<union> T)"
+ "\<lbrakk>connectedin X S; connectedin X T; \<not>separatedin X S T\<rbrakk> \<Longrightarrow> connectedin X (S \<union> 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)
--- 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 \<open>More properties of ARs, ANRs and ENRs\<close>
-lemma not_AR_empty [simp]: "~ AR({})"
+lemma not_AR_empty [simp]: "\<not> 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 \<in> components(- s)" and boc: "bounded c" and "a \<in> c"
- shows "~ (\<exists>g. continuous_on (s \<union> c) g \<and>
+ shows "\<not> (\<exists>g. continuous_on (s \<union> c) g \<and>
g ` (s \<union> c) \<subseteq> sphere 0 1 \<and>
(\<forall>x \<in> 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 \<notin> S"
shows "bounded (connected_component_set (- S) a) \<longleftrightarrow>
- ~(\<exists>c. homotopic_with (\<lambda>x. True) S (sphere 0 1)
+ \<not>(\<exists>c. homotopic_with (\<lambda>x. True) S (sphere 0 1)
(\<lambda>x. inverse(norm(x - a)) *\<^sub>R (x - a)) (\<lambda>x. c))"
(is "?lhs = ?rhs")
proof (cases "S = {}")
@@ -4550,9 +4550,9 @@
by (simp add: \<open>a \<notin> S\<close>)
obtain r where "r>0" and r: "S \<subseteq> ball 0 r"
using bounded_subset_ballD \<open>bounded S\<close> by blast
- have "~ ?rhs \<longleftrightarrow> ~ ?lhs"
+ have "\<not> ?rhs \<longleftrightarrow> \<not> ?lhs"
proof
- assume notr: "~ ?rhs"
+ assume notr: "\<not> ?rhs"
have nog: "\<nexists>g. continuous_on (S \<union> connected_component_set (- S) a) g \<and>
g ` (S \<union> connected_component_set (- S) a) \<subseteq> sphere 0 1 \<and>
(\<forall>x\<in>S. g x = (x - a) /\<^sub>R norm (x - a))"
@@ -4564,12 +4564,12 @@
using notr
by (auto simp: nullhomotopic_into_sphere_extension
[OF \<open>closed S\<close> continuous_on_Borsuk_map [OF \<open>a \<notin> S\<close>] False s01])
- with \<open>a \<notin> S\<close> show "~ ?lhs"
+ with \<open>a \<notin> S\<close> show "\<not> ?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 "\<not> ?lhs"
then obtain b where b: "b \<in> connected_component_set (- S) a" and "r \<le> norm b"
using bounded_iff linear by blast
then have bnot: "b \<notin> ball 0 r"
@@ -4593,7 +4593,7 @@
ultimately have "homotopic_with (\<lambda>x. True) S (sphere 0 1)
(\<lambda>x. (x - a) /\<^sub>R norm (x - a)) (\<lambda>x. c)"
by (meson homotopic_with_subset_left homotopic_with_trans r)
- then show "~ ?rhs"
+ then show "\<not> ?rhs"
by blast
qed
then show ?thesis by blast
--- 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 \<Rightarrow> real^'n"
- assumes net: "(f \<longlongrightarrow> l) net" "~(trivial_limit net)" and ev:"eventually (\<lambda>x. f(x)$i = b) net"
+ assumes net: "(f \<longlongrightarrow> l) net" "\<not> trivial_limit net" and ev:"eventually (\<lambda>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 "(\<exists>x. (x \<noteq> 0) \<and> A *v x = 0) \<longleftrightarrow> ~(rank A = CARD('n))"
+ shows "(\<exists>x. (x \<noteq> 0) \<and> A *v x = 0) \<longleftrightarrow> rank A \<noteq> CARD('n)"
by (meson matrix_left_invertible_injective full_rank_injective matrix_left_invertible_ker)
lemma%unimportant rank_eq_0: "rank A = 0 \<longleftrightarrow> A = 0" and rank_0 [simp]: "rank (0::real^'n^'m) = 0"
--- 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 "\<not> 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 \<gamma>) f l F"
and noleB: "\<And>t. t \<in> {0..1} \<Longrightarrow> norm (vector_derivative \<gamma> (at t)) \<le> B"
and \<gamma>: "valid_path \<gamma>"
- and [simp]: "~ (trivial_limit F)"
+ and [simp]: "\<not> trivial_limit F"
shows "l contour_integrable_on \<gamma>" "((\<lambda>n. contour_integral \<gamma> (f n)) \<longlongrightarrow> contour_integral \<gamma> l) F"
proof -
have "0 \<le> 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 "\<forall>\<^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 "\<not> trivial_limit F" "0 < r"
shows "l contour_integrable_on (circlepath z r)"
"((\<lambda>n. contour_integral (circlepath z r) (f n)) \<longlongrightarrow> 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 (\<lambda>n. continuous_on (cball z r) (f n) \<and> (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: "\<not> 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 (\<lambda>n. continuous_on (cball z r) (f n) \<and>
(\<forall>w \<in> 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: "\<not> trivial_limit F" and "0 < r"
obtains g' where
"continuous_on (cball z r) g"
"\<And>w. w \<in> ball z r \<Longrightarrow> (g has_field_derivative (g' w)) (at w) \<and> ((\<lambda>n. f' n w) \<longlongrightarrow> g' w) F"
--- 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 \<longlongrightarrow> l) F" and "~(trivial_limit F)" and "eventually P F" and "\<And>a. P a \<Longrightarrow> f a \<in> \<real>"
+ assumes "(f \<longlongrightarrow> l) F" and "\<not> trivial_limit F" and "eventually P F" and "\<And>a. P a \<Longrightarrow> f a \<in> \<real>"
shows "l \<in> \<real>"
proof (rule Lim_in_closed_set[OF closed_complex_Reals _ assms(2,1)])
show "eventually (\<lambda>x. f x \<in> \<real>) F"
--- 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) \<Longrightarrow> tan field_differentiable at z"
+lemma field_differentiable_at_tan: "cos z \<noteq> 0 \<Longrightarrow> 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 \<noteq> 0
\<Longrightarrow> 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) \<Longrightarrow> continuous (at z within s) tan"
+lemma continuous_within_tan: "cos z \<noteq> 0 \<Longrightarrow> continuous (at z within s) tan"
using continuous_at_imp_continuous_within isCont_tan by blast
-lemma continuous_on_tan [continuous_intros]: "(\<And>z. z \<in> s \<Longrightarrow> ~(cos z = 0)) \<Longrightarrow> continuous_on s tan"
+lemma continuous_on_tan [continuous_intros]: "(\<And>z. z \<in> s \<Longrightarrow> cos z \<noteq> 0) \<Longrightarrow> continuous_on s tan"
by (simp add: continuous_at_imp_continuous_on)
-lemma holomorphic_on_tan: "(\<And>z. z \<in> s \<Longrightarrow> ~(cos z = 0)) \<Longrightarrow> tan holomorphic_on s"
+lemma holomorphic_on_tan: "(\<And>z. z \<in> s \<Longrightarrow> cos z \<noteq> 0) \<Longrightarrow> tan holomorphic_on s"
by (simp add: field_differentiable_within_tan holomorphic_on_def)
@@ -1670,7 +1670,7 @@
lemma Ln_minus:
assumes "z \<noteq> 0"
- shows "Ln(-z) = (if Im(z) \<le> 0 \<and> ~(Re(z) < 0 \<and> Im(z) = 0)
+ shows "Ln(-z) = (if Im(z) \<le> 0 \<and> \<not>(Re(z) < 0 \<and> Im(z) = 0)
then Ln(z) + \<i> * pi
else Ln(z) - \<i> * pi)" (is "_ = ?rhs")
using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms
--- 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 \<subseteq> S"
- and fne: "~ f constant_on S"
+ and fne: "\<not> 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 \<subseteq> S"
- and fnc: "\<And>X. \<lbrakk>open X; X \<subseteq> S; X \<noteq> {}\<rbrakk> \<Longrightarrow> ~ f constant_on X"
+ and fnc: "\<And>X. \<lbrakk>open X; X \<subseteq> S; X \<noteq> {}\<rbrakk> \<Longrightarrow> \<not> f constant_on X"
shows "open (f ` U)"
proof -
have "S = \<Union>(components S)" by simp
@@ -518,7 +518,7 @@
assume "\<not> f constant_on S"
then have "open (f ` U)"
using open_mapping_thm assms by blast
- moreover have "~ open (f ` U)"
+ moreover have "\<not> open (f ` U)"
proof -
have "\<exists>t. cmod (f \<xi> - t) < e \<and> t \<notin> f ` U" if "0 < e" for e
apply (rule_tac x="if 0 < Re(f \<xi>) then f \<xi> + (e/2) else f \<xi> - (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 "\<xi> \<in> S" "f \<xi> = 0"
- and nonconst: "~ f constant_on S"
+ and nonconst: "\<not> f constant_on S"
obtains g r n
where "0 < n" "0 < r" "ball \<xi> r \<subseteq> S"
"g holomorphic_on ball \<xi> r"
@@ -1156,7 +1156,7 @@
obtain r where "r > 0" "ball \<xi> r \<subseteq> S" "inj_on f (ball \<xi> r)"
by (blast intro: that has_complex_derivative_locally_injective [OF assms])
then have \<xi>: "\<xi> \<in> ball \<xi> r" by simp
- then have nc: "~ f constant_on ball \<xi> r"
+ then have nc: "\<not> f constant_on ball \<xi> r"
using \<open>inj_on f (ball \<xi> r)\<close> injective_not_constant by fastforce
have holf': "f holomorphic_on ball \<xi> r"
using \<open>ball \<xi> r \<subseteq> S\<close> 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 \<subseteq> s \<Longrightarrow> ball a t \<inter> frontier s \<noteq> {}"
+ have "\<not> ball a t \<subseteq> s \<Longrightarrow> ball a t \<inter> frontier s \<noteq> {}"
apply (rule connected_Int_frontier [of "ball a t" s], simp_all)
using \<open>0 < t\<close> \<open>a \<in> s\<close> centre_in_ball apply blast
done
--- 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:
"\<lbrakk>closedin (subtopology euclidean (s \<union> t)) s; closedin (subtopology euclidean (s \<union> t)) t;
continuous_on s f; continuous_on t g;
- \<And>x. \<lbrakk>x \<in> s \<and> ~P x \<or> x \<in> t \<and> P x\<rbrakk> \<Longrightarrow> f x = g x\<rbrakk>
+ \<And>x. \<lbrakk>x \<in> s \<and> \<not>P x \<or> x \<in> t \<and> P x\<rbrakk> \<Longrightarrow> f x = g x\<rbrakk>
\<Longrightarrow> continuous_on (s \<union> t) (\<lambda>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 \<union> T)"
proof -
have *: "\<lbrakk>\<And>x y. P x y \<longleftrightarrow> P y x; \<And>x y. P x y \<Longrightarrow> S \<subseteq> x \<or> S \<subseteq> y;
- \<And>x y. \<lbrakk>P x y; S \<subseteq> x\<rbrakk> \<Longrightarrow> False\<rbrakk> \<Longrightarrow> ~(\<exists>x y. (P x y))" for P
+ \<And>x y. \<lbrakk>P x y; S \<subseteq> x\<rbrakk> \<Longrightarrow> False\<rbrakk> \<Longrightarrow> \<not>(\<exists>x y. (P x y))" for P
by metis
show ?thesis
unfolding connected_closedin_eq
--- 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 \<open>Continuous extensions of functions: Urysohn's lemma, Dugundji extension theorem, Tietze\<close>
+section \<open>Continuous Extensions of Functions\<close>
theory Continuous_Extension
imports Starlike
@@ -112,7 +112,9 @@
qed
-subsection\<open>Urysohn's lemma (for Euclidean spaces, where the proof is easy using distances)\<close>
+subsection\<open>Urysohn's Lemma for Euclidean Spaces\<close>
+
+text \<open>For Euclidean spaces the proof is easy using distances.\<close>
lemma Urysohn_both_ne:
assumes US: "closedin (subtopology euclidean U) S"
@@ -130,7 +132,7 @@
using \<open>T \<noteq> {}\<close> UT setdist_eq_0_closedin by auto
have sdpos: "0 < setdist {x} S + setdist {x} T" if "x \<in> U" for x
proof -
- have "~ (setdist {x} S = 0 \<and> setdist {x} T = 0)"
+ have "\<not> (setdist {x} S = 0 \<and> 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\<open>The Dugundji extension theorem and Tietze variants as corollaries\<close>
+subsection\<open>The Dugundji Extension Theorem and Tietze Variants\<close>
text%important\<open>J. Dugundji. An extension of Tietze's theorem. Pacific J. Math. Volume 1, Number 3 (1951), 353-367.
https://projecteuclid.org/euclid.pjm/1103052106\<close>
--- 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 \<open>Some explicit formulations (from Lars Schewe)\<close>
+subsubsection%unimportant \<open>Some explicit formulations\<close>
+
+text "Formalized by Lars Schewe."
lemma affine:
fixes V::"'a::real_vector set"
@@ -2070,7 +2072,9 @@
qed
-subsection \<open>Affine dependence and consequential theorems (from Lars Schewe)\<close>
+subsection \<open>Affine dependence and consequential theorems\<close>
+
+text "Formalized by Lars Schewe."
definition%important affine_dependent :: "'a::real_vector set \<Rightarrow> bool"
where "affine_dependent s \<longleftrightarrow> (\<exists>x\<in>s. x \<in> affine hull (s - {x}))"
@@ -2082,11 +2086,11 @@
done
lemma affine_independent_subset:
- shows "\<lbrakk>~ affine_dependent t; s \<subseteq> t\<rbrakk> \<Longrightarrow> ~ affine_dependent s"
+ shows "\<lbrakk>\<not> affine_dependent t; s \<subseteq> t\<rbrakk> \<Longrightarrow> \<not> affine_dependent s"
by (metis affine_dependent_subset)
lemma affine_independent_Diff:
- "~ affine_dependent s \<Longrightarrow> ~ affine_dependent(s - t)"
+ "\<not> affine_dependent s \<Longrightarrow> \<not> affine_dependent(s - t)"
by (meson Diff_subset affine_dependent_subset)
proposition affine_dependent_explicit:
@@ -2601,7 +2605,9 @@
qed (use assms in \<open>auto simp: convex_explicit\<close>)
-subsubsection%unimportant \<open>Another formulation from Lars Schewe\<close>
+subsubsection%unimportant \<open>Another formulation\<close>
+
+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 \<longleftrightarrow> finite s \<and> aff_dim s = int(card s) - 1"
+ shows "\<not> affine_dependent s \<longleftrightarrow> finite s \<and> 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 \<in> S"
+ assumes "\<not> affine_dependent S" "a \<in> S"
shows "card S = dim {x - a|x. x \<in> S} + 1"
proof -
have 1: "{b - a|b. b \<in> (S - {a})} \<subseteq> {x - a|x. x \<in> S}" by auto
@@ -3766,7 +3772,7 @@
show ?thesis
proof safe
assume 0: "aff_dim S = 0"
- have "~ {a,b} \<subseteq> S" if "b \<noteq> a" for b
+ have "\<not> {a,b} \<subseteq> S" if "b \<noteq> a" for b
by (metis "0" aff_dim_2 aff_dim_subset not_one_le_zero that)
then show "\<exists>a. S = {a}"
using \<open>a \<in> S\<close> by blast
@@ -3798,7 +3804,7 @@
lemma disjoint_affine_hull:
fixes s :: "'n::euclidean_space set"
- assumes "~ affine_dependent s" "t \<subseteq> s" "u \<subseteq> s" "t \<inter> u = {}"
+ assumes "\<not> affine_dependent s" "t \<subseteq> s" "u \<subseteq> s" "t \<inter> u = {}"
shows "(affine hull t) \<inter> (affine hull u) = {}"
proof -
have "finite s" using assms by (simp add: aff_independent_finite)
@@ -3813,7 +3819,7 @@
have [simp]: "s \<inter> t = t" "s \<inter> - t \<inter> u = u" using assms by auto
have "sum c s = 0"
by (simp add: c_def comm_monoid_add_class.sum.If_cases \<open>finite s\<close> sum_negf)
- moreover have "~ (\<forall>v\<in>s. c v = 0)"
+ moreover have "\<not> (\<forall>v\<in>s. c v = 0)"
by (metis (no_types) IntD1 \<open>s \<inter> t = t\<close> a1 c_def sum_not_0 zero_neq_one)
moreover have "(\<Sum>v\<in>s. c v *\<^sub>R v) = 0"
by (simp add: c_def if_smult sum_negf
@@ -5747,7 +5753,9 @@
qed auto
-subsection \<open>Radon's theorem (from Lars Schewe)\<close>
+subsection \<open>Radon's theorem\<close>
+
+text "Formalized by Lars Schewe."
lemma radon_ex_lemma:
assumes "finite c" "affine_dependent c"
--- 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 \<noteq> 0 \<Longrightarrow> ~(u \<times> axis 1 1 = 0) \<or> ~(u \<times> axis 2 1 = 0) \<or> ~(u \<times> axis 3 1 = 0)"
+ "u \<noteq> 0 \<Longrightarrow> u \<times> axis 1 1 \<noteq> 0 \<or> u \<times> axis 2 1 \<noteq> 0 \<or> u \<times> axis 3 1 \<noteq> 0"
by (clarsimp simp add: axis_def cross3_simps) (metis vector_3 exhaust_3)
lemma cross_dot_cancel:
--- 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 \<open>closed S\<close>, of a])
fix c x
assume cx: "a + c *\<^sub>R x \<in> S" "0 \<le> c" "a + x \<in> S"
- with star have "~ (c < 1)" by auto
- moreover have "~ (c > 1)"
+ with star have "\<not> (c < 1)" by auto
+ moreover have "\<not> (c > 1)"
using star [of "1/c" "c *\<^sub>R x"] cx by force
ultimately show "c = 1" by arith
qed
--- 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 \<open>~ s \<le> 0\<close> have "r > 0" "s > 0" by auto
+ with \<open>\<not> s \<le> 0\<close> have "r > 0" "s > 0" by auto
show ?thesis
apply (rule inessential_spheremap_lowdim_gen [of "cball a r" "cball b s" f])
using \<open>0 < r\<close> \<open>0 < s\<close> assms(1)
@@ -414,7 +414,7 @@
assumes fin: "finite \<F>"
and "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>g. continuous_on S g \<and> g ` S \<subseteq> T \<and> (\<forall>x \<in> S \<inter> K. g x = h x)"
and "\<And>S. S \<in> \<F> \<Longrightarrow> closed S"
- and K: "\<And>X Y. \<lbrakk>X \<in> \<F>; Y \<in> \<F>; ~ X \<subseteq> Y; ~ Y \<subseteq> X\<rbrakk> \<Longrightarrow> X \<inter> Y \<subseteq> K"
+ and K: "\<And>X Y. \<lbrakk>X \<in> \<F>; Y \<in> \<F>; \<not> X \<subseteq> Y; \<not> Y \<subseteq> X\<rbrakk> \<Longrightarrow> X \<inter> Y \<subseteq> K"
shows "\<exists>g. continuous_on (\<Union>\<F>) g \<and> g ` (\<Union>\<F>) \<subseteq> T \<and> (\<forall>x \<in> \<Union>\<F> \<inter> 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 \<in> \<G> \<union> ?Faces" for S
using that \<open>\<G> \<subseteq> \<F>\<close> face_of_polytope_polytope poly polytope_imp_closed by blast
have K: "X \<inter> Y \<subseteq> \<Union>(\<G> \<union> {D. \<exists>C\<in>\<F>. D face_of C \<and> aff_dim D < int p})"
- if "X \<in> \<G> \<union> ?Faces" "Y \<in> \<G> \<union> ?Faces" "~ Y \<subseteq> X" for X Y
+ if "X \<in> \<G> \<union> ?Faces" "Y \<in> \<G> \<union> ?Faces" "\<not> Y \<subseteq> X" for X Y
proof -
have ff: "X \<inter> Y face_of X \<and> X \<inter> Y face_of Y"
if XY: "X face_of D" "Y face_of E" and DE: "D \<in> \<F>" "E \<in> \<F>" for D E
@@ -669,7 +669,7 @@
assumes "finite \<F>"
and \<F>: "\<And>X. X \<in> \<F> \<Longrightarrow> \<exists>a g. a \<notin> U \<and> continuous_on (X - {a}) g \<and> g ` (X - {a}) \<subseteq> T \<and> (\<forall>x \<in> X \<inter> K. g x = h x)"
and clo: "\<And>X. X \<in> \<F> \<Longrightarrow> closed X"
- and K: "\<And>X Y. \<lbrakk>X \<in> \<F>; Y \<in> \<F>; ~(X \<subseteq> Y); ~(Y \<subseteq> X)\<rbrakk> \<Longrightarrow> X \<inter> Y \<subseteq> K"
+ and K: "\<And>X Y. \<lbrakk>X \<in> \<F>; Y \<in> \<F>; \<not> X \<subseteq> Y; \<not> Y \<subseteq> X\<rbrakk> \<Longrightarrow> X \<inter> Y \<subseteq> K"
obtains C g where "finite C" "disjnt C U" "card C \<le> card \<F>" "continuous_on (\<Union>\<F> - C) g"
"g ` (\<Union>\<F> - C) \<subseteq> T"
"\<And>x. x \<in> (\<Union>\<F> - C) \<inter> K \<Longrightarrow> g x = h x"
@@ -1188,8 +1188,8 @@
then have "ball a d \<inter> U \<subseteq> C"
by auto
obtain h k where homhk: "homeomorphism (S \<union> C) (S \<union> C) h k"
- and subC: "{x. (~ (h x = x \<and> k x = x))} \<subseteq> C"
- and bou: "bounded {x. (~ (h x = x \<and> k x = x))}"
+ and subC: "{x. (\<not> (h x = x \<and> k x = x))} \<subseteq> C"
+ and bou: "bounded {x. (\<not> (h x = x \<and> k x = x))}"
and hin: "\<And>x. x \<in> C \<inter> K \<Longrightarrow> h x \<in> ball a d \<inter> U"
proof (rule homeomorphism_grouping_points_exists_gen [of C "ball a d \<inter> U" "C \<inter> K" "S \<union> C"])
show "openin (subtopology euclidean C) (ball a d \<inter> U)"
@@ -1539,7 +1539,7 @@
by (simp add: assms compact_imp_bounded)
then obtain b where b: "S \<subseteq> cbox (-b) b"
using bounded_subset_cbox_symmetric by blast
- define LU where "LU \<equiv> L \<union> (\<Union> {C \<in> components (T - S). ~bounded C} - cbox (-(b+One)) (b+One))"
+ define LU where "LU \<equiv> L \<union> (\<Union> {C \<in> components (T - S). \<not>bounded C} - cbox (-(b+One)) (b+One))"
obtain K g where "finite K" "K \<subseteq> LU" "K \<subseteq> T" "disjnt K S"
and contg: "continuous_on (T - K) g"
and gim: "g ` (T - K) \<subseteq> rel_frontier U"
@@ -1699,7 +1699,7 @@
theorem%important Borsuk_separation_theorem_gen:
fixes S :: "'a::euclidean_space set"
assumes "compact S"
- shows "(\<forall>c \<in> components(- S). ~bounded c) \<longleftrightarrow>
+ shows "(\<forall>c \<in> components(- S). \<not>bounded c) \<longleftrightarrow>
(\<forall>f. continuous_on S f \<and> f ` S \<subseteq> sphere (0::'a) 1
\<longrightarrow> (\<exists>c. homotopic_with (\<lambda>x. True) S (sphere 0 1) f (\<lambda>x. c)))"
(is "?lhs = ?rhs")
@@ -5204,7 +5204,7 @@
proof (rule ccontr)
let ?S = "S \<union> \<Union>{C \<in> components(- (S \<union> T)). frontier C \<subseteq> S}"
let ?T = "T \<union> \<Union>{C \<in> components(- (S \<union> T)). frontier C \<subseteq> T}"
- assume "~ thesis"
+ assume "\<not> thesis"
with that have *: "frontier C \<inter> S = {} \<or> frontier C \<inter> T = {}"
if C: "C \<in> components (- (S \<union> T))" "C \<noteq> {}" for C
using C by blast
--- 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 (\<F> n x / (x - y) - \<F> n x / (x - z)) \<le> e / r"
if "cmod (x - z) = r/3 + r/3" for x
proof -
- have "~ (cmod (x - y) < r/3)"
+ have "\<not> (cmod (x - y) < r/3)"
using y_near_z(1) that \<open>M > 0\<close> \<open>r > 0\<close>
by (metis (full_types) norm_diff_triangle_less norm_minus_commute order_less_irrefl)
then have r4_le_xy: "r/4 \<le> cmod (x - y)"
@@ -1000,7 +1000,7 @@
and holf: "\<And>n::nat. \<F> n holomorphic_on S"
and holg: "g holomorphic_on S"
and ul_g: "\<And>K. \<lbrakk>compact K; K \<subseteq> S\<rbrakk> \<Longrightarrow> uniform_limit K \<F> g sequentially"
- and nonconst: "~ g constant_on S"
+ and nonconst: "\<not> g constant_on S"
and nz: "\<And>n z. z \<in> S \<Longrightarrow> \<F> n z \<noteq> 0"
and "z0 \<in> S"
shows "g z0 \<noteq> 0"
@@ -1164,7 +1164,7 @@
and holf: "\<And>n::nat. \<F> n holomorphic_on S"
and holg: "g holomorphic_on S"
and ul_g: "\<And>K. \<lbrakk>compact K; K \<subseteq> S\<rbrakk> \<Longrightarrow> uniform_limit K \<F> g sequentially"
- and nonconst: "~ g constant_on S"
+ and nonconst: "\<not> g constant_on S"
and inj: "\<And>n. inj_on (\<F> 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 \<in> S"
- have "~ (\<forall>r>0. \<exists>h\<in>Y. r < cmod (h v))"
+ have "\<not> (\<forall>r>0. \<exists>h\<in>Y. r < cmod (h v))"
proof
assume "\<forall>r>0. \<exists>h\<in>Y. r < cmod (h v)"
then have "\<forall>n. \<exists>h\<in>Y. Suc n < cmod (h v)"
--- 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 \<le> content X"
by simp
-corollary content_nonneg [simp]: "~ content (cbox a b) < 0"
+corollary content_nonneg [simp]: "\<not> content (cbox a b) < 0"
using not_le by blast
lemma content_pos_lt: "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i \<Longrightarrow> 0 < content (cbox a b)"
@@ -287,12 +287,12 @@
definition integrable_on (infixr "integrable'_on" 46)
where "f integrable_on i \<longleftrightarrow> (\<exists>y. (f has_integral y) i)"
-definition "integral i f = (SOME y. (f has_integral y) i \<or> ~ f integrable_on i \<and> y=0)"
+definition "integral i f = (SOME y. (f has_integral y) i \<or> \<not> f integrable_on i \<and> y=0)"
lemma integrable_integral[intro]: "f integrable_on i \<Longrightarrow> (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 \<Longrightarrow> integral i f = 0"
+lemma not_integrable_integral: "\<not> f integrable_on i \<Longrightarrow> integral i f = 0"
unfolding integrable_on_def integral_def by blast
lemma has_integral_integrable[dest]: "(f has_integral i) s \<Longrightarrow> f integrable_on s"
@@ -356,7 +356,7 @@
lemma has_integral_iff: "(f has_integral i) S \<longleftrightarrow> (f integrable_on S \<and> integral S f = i)"
by blast
-lemma eq_integralD: "integral k f = y \<Longrightarrow> (f has_integral y) k \<or> ~ f integrable_on k \<and> y=0"
+lemma eq_integralD: "integral k f = y \<Longrightarrow> (f has_integral y) k \<or> \<not> f integrable_on k \<and> 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 "~ (\<lambda>x. f x * c) integrable_on S"
+ case False then have "\<not> (\<lambda>x. f x * c) integrable_on S"
using has_integral_mult_left [of "(\<lambda>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 "~ (\<lambda>x. c *\<^sub>R f x) integrable_on S"
+ case False then have "\<not> (\<lambda>x. c *\<^sub>R f x) integrable_on S"
using has_integral_cmul [of "(\<lambda>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 "~ (\<lambda>x. - f x) integrable_on S"
+ case False then have "\<not> (\<lambda>x. - f x) integrable_on S"
using has_integral_neg [of "(\<lambda>x. - f x)" _ S ] by auto
with False show ?thesis by (simp add: not_integrable_integral)
qed
--- 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 \<in> rel_interior S"
- and y: "y \<in> S" and xy: "~(x = y \<and> S = {x})"
+ and y: "y \<in> S" and xy: "\<not>(x = y \<and> S = {x})"
obtains z where "z \<in> rel_frontier S" "y \<in> closed_segment x z"
"open_segment x z \<subseteq> rel_interior S"
proof (cases "x=y")
--- 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 (\<Sum>(x,K) \<in> T'. ?CI K h x - ?CI K f x)
\<le> 2*\<epsilon>/3"
proof -
- define T'' where "T'' \<equiv> {(x,K) \<in> T'. ~ (K \<subseteq> {x. x \<bullet> i \<le> c})}"
+ define T'' where "T'' \<equiv> {(x,K) \<in> T'. \<not> (K \<subseteq> {x. x \<bullet> i \<le> c})}"
then have "T'' \<subseteq> T'"
by auto
then have "finite T''"
--- 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 \<union> path_image d \<union> path_image g)) x y"
by (simp add: Un_ac)
- moreover have "~(connected_component (- (path_image c)) x y)"
+ moreover have "\<not>(connected_component (- (path_image c)) x y)"
by (metis (no_types, lifting) \<open>\<not> bounded outer\<close> \<open>bounded inner\<close> \<open>x \<in> inner\<close> \<open>y \<in> outer\<close> componentsE connected_component_eq inner mem_Collect_eq outer)
ultimately show False
by (auto simp: ud_Un [symmetric] connected_component_def)
--- 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) \<Longrightarrow> {a..b} = {}"
+ "(\<not> a \<le> b) \<Longrightarrow> {a..b} = {}"
by (auto)
instance real :: ordered_euclidean_space
--- 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 \<open>d1 > 0\<close> \<open>d2 > 0\<close> by (simp add: min_def)
have dist2: "dist (1 / 2 + min (1 / 2) (min d1 d2) / 4) (1 / 2) < d2"
using \<open>d1 > 0\<close> \<open>d2 > 0\<close> by (simp add: min_def dist_norm)
- have [simp]: "~ min (1 / 2) (min d1 d2) \<le> 0"
+ have [simp]: "\<not> min (1 / 2) (min d1 d2) \<le> 0"
using \<open>d1 > 0\<close> \<open>d2 > 0\<close> 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 \<le> u" "u \<le> 1" for u
using * [of 0 "(u + 1) / 2"] that assms arc_distinct_ends [OF \<open>?lhs\<close>]
by (auto simp: joinpaths_def pathstart_def pathfinish_def split_ifs divide_simps)
- then have n1: "~ (pathstart g1 \<in> path_image g2)"
+ then have n1: "pathstart g1 \<notin> path_image g2"
unfolding pathstart_def path_image_def
using atLeastAtMost_iff by blast
show ?rhs using \<open>?lhs\<close>
@@ -2442,7 +2442,7 @@
lemma cobounded_unbounded_component:
fixes s :: "'a :: euclidean_space set"
assumes "bounded (-s)"
- shows "\<exists>x. x \<in> s \<and> ~ bounded (connected_component_set s x)"
+ shows "\<exists>x. x \<in> s \<and> \<not> bounded (connected_component_set s x)"
proof -
obtain i::'a where i: "i \<in> Basis"
using nonempty_Basis by blast
@@ -2450,7 +2450,7 @@
using bounded_subset_ballD [OF assms, of 0] by auto
then have *: "\<And>x. B \<le> norm x \<Longrightarrow> x \<in> s"
by (force simp: ball_def dist_norm)
- have unbounded_inner: "~ bounded {x. inner i x \<ge> B}"
+ have unbounded_inner: "\<not> bounded {x. inner i x \<ge> B}"
apply (auto simp: bounded_def dist_norm)
apply (rule_tac x="x + (max B e + 1 + \<bar>i \<bullet> x\<bar>) *\<^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 \<le> DIM('a)"
- and bo: "~ bounded(connected_component_set s x)"
- "~ bounded(connected_component_set s y)"
+ and bo: "\<not> bounded(connected_component_set s x)"
+ "\<not> bounded(connected_component_set s y)"
shows "connected_component_set s x = connected_component_set s y"
proof -
obtain i::'a where i: "i \<in> Basis"
@@ -2507,7 +2507,7 @@
lemma cobounded_unbounded_components:
fixes s :: "'a :: euclidean_space set"
- shows "bounded (-s) \<Longrightarrow> \<exists>c. c \<in> components s \<and> ~bounded c"
+ shows "bounded (-s) \<Longrightarrow> \<exists>c. c \<in> components s \<and> \<not>bounded c"
by (metis cobounded_unbounded_component components_def imageI)
lemma cobounded_unique_unbounded_components:
@@ -2532,9 +2532,9 @@
"inside S \<equiv> {x. (x \<notin> S) \<and> bounded(connected_component_set ( - S) x)}"
definition%important outside where
- "outside S \<equiv> -S \<inter> {x. ~ bounded(connected_component_set (- S) x)}"
-
-lemma outside: "outside S = {x. ~ bounded(connected_component_set (- S) x)}"
+ "outside S \<equiv> -S \<inter> {x. \<not> bounded(connected_component_set (- S) x)}"
+
+lemma outside: "outside S = {x. \<not> 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 \<inter> S = {}"
@@ -2619,7 +2619,7 @@
lemma unbounded_outside:
fixes S :: "'a::{real_normed_vector, perfect_space} set"
- shows "bounded S \<Longrightarrow> ~ bounded(outside S)"
+ shows "bounded S \<Longrightarrow> \<not> 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 \<le> DIM('a)"
- shows "- (outside S) = {x. \<forall>B. \<exists>y. B < norm(y) \<and> ~ (connected_component (- S) x y)}"
+ shows "- (outside S) = {x. \<forall>B. \<exists>y. B < norm(y) \<and> \<not> connected_component (- S) x y}"
proof -
obtain B::real where B: "0 < B" and Bno: "\<And>x. x \<in> S \<Longrightarrow> norm x \<le> 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 \<le> DIM('a)"
- shows "- (outside S) = {x. \<forall>B. \<exists>y. B \<le> norm(y) \<and> ~ (connected_component (- S) x y)}"
+ shows "- (outside S) = {x. \<forall>B. \<exists>y. B \<le> norm(y) \<and> \<not> 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 \<le> DIM('a)"
- shows "inside S = {x. (x \<notin> S) \<and> (\<forall>B. \<exists>y. B < norm(y) \<and> ~(connected_component (- S) x y))}"
+ shows "inside S = {x. (x \<notin> S) \<and> (\<forall>B. \<exists>y. B < norm(y) \<and> \<not> 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 \<le> DIM('a)"
- shows "inside S = {x. (x \<notin> S) \<and> (\<forall>B. \<exists>y. B \<le> norm(y) \<and> ~(connected_component (- S) x y))}"
+ shows "inside S = {x. (x \<notin> S) \<and> (\<forall>B. \<exists>y. B \<le> norm(y) \<and> \<not> 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 \<union> U = - S"
+ assumes "connected U" and "\<not> bounded U" and "T \<union> U = - S"
shows "inside S \<subseteq> 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: "\<not> connected(- S)"
and C: "C \<in> components (- S)"
and conn: "\<And>T. \<lbrakk>closed T; T \<subset> S\<rbrakk> \<Longrightarrow> connected(- T)"
shows "frontier C = S"
@@ -3319,7 +3319,7 @@
lemma bounded_unique_outside:
fixes s :: "'a :: euclidean_space set"
- shows "\<lbrakk>bounded s; DIM('a) \<ge> 2\<rbrakk> \<Longrightarrow> (c \<in> components (- s) \<and> ~bounded c \<longleftrightarrow> c = outside s)"
+ shows "\<lbrakk>bounded s; DIM('a) \<ge> 2\<rbrakk> \<Longrightarrow> (c \<in> components (- s) \<and> \<not> bounded c \<longleftrightarrow> 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 \<inter> (- S) = {}) \<and> ~(b - (- S) = {}) \<Longrightarrow> (b \<inter> f \<noteq> {}))
+ have **: "(b \<inter> (- S) \<noteq> {} \<and> b - (- S) \<noteq> {} \<Longrightarrow> b \<inter> f \<noteq> {})
\<Longrightarrow> (b \<inter> S \<noteq> {}) \<Longrightarrow> b \<inter> f = {} \<Longrightarrow>
b \<subseteq> S" for b f and S :: "'b set"
by blast
@@ -5065,7 +5065,7 @@
done
qed
-subsection\<open>Sort of induction principle for connected sets\<close>
+subsection\<open>An induction principle for connected sets\<close>
proposition connected_induction:
assumes "connected S"
@@ -5084,7 +5084,7 @@
done
have 2: "openin (subtopology euclidean S)
{b. \<exists>T. openin (subtopology euclidean S) T \<and>
- b \<in> T \<and> (\<forall>x\<in>T. P x \<longrightarrow> ~ Q x)}"
+ b \<in> T \<and> (\<forall>x\<in>T. P x \<longrightarrow> \<not> Q x)}"
apply (subst openin_subopen, clarify)
apply (rule_tac x=T in exI, auto)
done
@@ -5742,7 +5742,7 @@
qed blast
-subsection\<open>Important special cases of local connectedness and path connectedness\<close>
+subsection\<open>Special cases of local connectedness and path connectedness\<close>
lemma locally_connected_1:
assumes
@@ -7155,7 +7155,7 @@
lemma connected_card_eq_iff_nontrivial:
fixes S :: "'a::metric_space set"
- shows "connected S \<Longrightarrow> uncountable S \<longleftrightarrow> ~(\<exists>a. S \<subseteq> {a})"
+ shows "connected S \<Longrightarrow> uncountable S \<longleftrightarrow> \<not>(\<exists>a. S \<subseteq> {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" "\<not> 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 \<in> U"
using \<open>a \<in> U\<close> \<open>b \<in> U\<close> \<open>convex U\<close> convex_contains_segment by force
obtain c where "c \<in> U" and nc_abc: "\<not> collinear {a,b,c}"
- by (metis False \<open>a \<in> U\<close> \<open>b \<in> U\<close> \<open>~ collinear U\<close> collinear_triples insert_absorb)
+ by (metis False \<open>a \<in> U\<close> \<open>b \<in> U\<close> \<open>\<not> collinear U\<close> collinear_triples insert_absorb)
have ncoll_mca: "\<not> collinear {?m,c,a}"
by (metis (full_types) \<open>a \<noteq> ?m\<close> collinear_3_trans collinear_midpoint insert_commute nc_abc)
have ncoll_mcb: "\<not> 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" "\<not> 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 "\<not> 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 \<open>x \<in> U\<close> have x: "x \<in> ball x r \<inter> affine hull S"
by auto
- have "~ S \<subseteq> {x}"
- using \<open>~ collinear S\<close> collinear_subset by blast
+ have "\<not> S \<subseteq> {x}"
+ using \<open>\<not> collinear S\<close> collinear_subset by blast
then obtain x' where "x' \<noteq> x" "x' \<in> S"
by blast
obtain y where y: "y \<noteq> x" "y \<in> ball x r \<inter> 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 "\<not> 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\<open>Self-homeomorphisms shuffling points about in various ways\<close>
+subsection\<open>Self-homeomorphisms shuffling points about\<close>
subsubsection%unimportant\<open>The theorem \<open>homeomorphism_moving_points_exists\<close>\<close>
@@ -7610,7 +7610,7 @@
assumes "affine T" "a \<in> T" and ST: "ball a r \<inter> T \<subseteq> S" "S \<subseteq> T"
and u: "u \<in> ball a r \<inter> T" and v: "v \<in> ball a r \<inter> T"
obtains f g where "homeomorphism S S f g"
- "f u = v" "{x. ~ (f x = x \<and> g x = x)} \<subseteq> ball a r \<inter> T"
+ "f u = v" "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> ball a r \<inter> T"
proof -
obtain f g where hom: "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f g"
and "f u = v" and fid: "\<And>x. \<lbrakk>x \<in> sphere a r; x \<in> T\<rbrakk> \<Longrightarrow> f x = x"
@@ -7690,14 +7690,14 @@
and TS: "T \<subseteq> affine hull S"
and S: "connected S" "a \<in> S" "b \<in> S"
obtains f g where "homeomorphism T T f g" "f a = b"
- "{x. ~ (f x = x \<and> g x = x)} \<subseteq> S"
- "bounded {x. ~ (f x = x \<and> g x = x)}"
+ "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S"
+ "bounded {x. \<not> (f x = x \<and> g x = x)}"
proof -
have 1: "\<exists>h k. homeomorphism T T h k \<and> h (f d) = d \<and>
- {x. ~ (h x = x \<and> k x = x)} \<subseteq> S \<and> bounded {x. ~ (h x = x \<and> k x = x)}"
+ {x. \<not> (h x = x \<and> k x = x)} \<subseteq> S \<and> bounded {x. \<not> (h x = x \<and> k x = x)}"
if "d \<in> S" "f d \<in> S" and homfg: "homeomorphism T T f g"
- and S: "{x. ~ (f x = x \<and> g x = x)} \<subseteq> S"
- and bo: "bounded {x. ~ (f x = x \<and> g x = x)}" for d f g
+ and S: "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S"
+ and bo: "bounded {x. \<not> (f x = x \<and> 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. \<not> ((f2 \<circ> f1) x = x \<and> (g1 \<circ> g2) x = x)} \<subseteq> S"
using sub by force
- have "bounded ({x. ~(f1 x = x \<and> g1 x = x)} \<union> {x. ~(f2 x = x \<and> g2 x = x)})"
+ have "bounded ({x. \<not>(f1 x = x \<and> g1 x = x)} \<union> {x. \<not>(f2 x = x \<and> g2 x = x)})"
using bo by simp
then show "bounded {x. \<not> ((f2 \<circ> f1) x = x \<and> (g1 \<circ> g2) x = x)}"
by (rule bounded_subset) auto
@@ -7753,7 +7753,7 @@
done
qed
have "\<exists>f g. homeomorphism T T f g \<and> f a = b \<and>
- {x. ~ (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. ~ (f x = x \<and> g x = x)}"
+ {x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. \<not> (f x = x \<and> 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 \<subseteq> T" "T \<subseteq> affine hull S" "connected S"
shows "\<exists>f g. homeomorphism T T f g \<and> (\<forall>i \<in> K. f(x i) = y i) \<and>
- {x. ~ (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. ~ (f x = x \<and> g x = x)}"
+ {x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. \<not> (f x = x \<and> g x = x)}"
using assms
proof (induction K)
case empty
@@ -7783,11 +7783,11 @@
and xyS: "\<And>i. i \<in> K \<Longrightarrow> x i \<in> S \<and> y i \<in> S"
by (simp_all add: pairwise_insert)
obtain f g where homfg: "homeomorphism T T f g" and feq: "\<And>i. i \<in> K \<Longrightarrow> f(x i) = y i"
- and fg_sub: "{x. ~ (f x = x \<and> g x = x)} \<subseteq> S"
- and bo_fg: "bounded {x. ~ (f x = x \<and> g x = x)}"
+ and fg_sub: "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S"
+ and bo_fg: "bounded {x. \<not> (f x = x \<and> g x = x)}"
using insert.IH [OF xyS pw] insert.prems by (blast intro: that)
then have "\<exists>f g. homeomorphism T T f g \<and> (\<forall>i \<in> K. f(x i) = y i) \<and>
- {x. ~ (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. ~ (f x = x \<and> g x = x)}"
+ {x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. \<not> (f x = x \<and> 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. \<not> ((h \<circ> f) x = x \<and> (g \<circ> k) x = x)} \<subseteq> S"
using fg_sub hk_sub by force
- have "bounded ({x. ~(f x = x \<and> g x = x)} \<union> {x. ~(h x = x \<and> k x = x)})"
+ have "bounded ({x. \<not>(f x = x \<and> g x = x)} \<union> {x. \<not>(h x = x \<and> k x = x)})"
using bo_fg bo_hk bounded_Un by blast
then show "bounded {x. \<not> ((h \<circ> f) x = x \<and> (g \<circ> k) x = x)}"
by (rule bounded_subset) auto
@@ -7846,7 +7846,7 @@
and pw: "pairwise (\<lambda>i j. (x i \<noteq> x j) \<and> (y i \<noteq> y j)) K"
and S: "S \<subseteq> T" "T \<subseteq> affine hull S" "connected S"
obtains f g where "homeomorphism T T f g" "\<And>i. i \<in> K \<Longrightarrow> f(x i) = y i"
- "{x. ~ (f x = x \<and> g x = x)} \<subseteq> S" "bounded {x. (~ (f x = x \<and> g x = x))}"
+ "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S" "bounded {x. (\<not> (f x = x \<and> 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 \<noteq> {}" "finite K" "K \<subseteq> S" "U \<subseteq> S" "S \<subseteq> T"
obtains f g where "homeomorphism T T f g"
- "\<And>x. x \<in> K \<Longrightarrow> f x \<in> U" "{x. (~ (f x = x \<and> g x = x))} \<subseteq> S"
- "bounded {x. (~ (f x = x \<and> g x = x))}"
+ "\<And>x. x \<in> K \<Longrightarrow> f x \<in> U" "{x. (\<not> (f x = x \<and> g x = x))} \<subseteq> S"
+ "bounded {x. (\<not> (f x = x \<and> g x = x))}"
proof -
obtain c d where "box c d \<noteq> {}" "cbox c d \<subseteq> 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 \<noteq> {}" "finite K" "K \<subseteq> S" "U \<subseteq> S" "S \<subseteq> T"
- obtains f g where "homeomorphism T T f g" "{x. (~ (f x = x \<and> g x = x))} \<subseteq> S"
- "bounded {x. (~ (f x = x \<and> g x = x))}" "\<And>x. x \<in> K \<Longrightarrow> f x \<in> U"
+ obtains f g where "homeomorphism T T f g" "{x. (\<not> (f x = x \<and> g x = x))} \<subseteq> S"
+ "bounded {x. (\<not> (f x = x \<and> g x = x))}" "\<And>x. x \<in> K \<Longrightarrow> f x \<in> U"
proof (cases "2 \<le> DIM('a)")
case True
have TS: "T \<subseteq> 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 \<and> g x = x))})"
+ have "bounded (j ` {x. (\<not> (f x = x \<and> g x = x))})"
by (rule bounded_linear_image [OF bou]) (use \<open>linear j\<close> linear_conv_bounded_linear in auto)
moreover
- have *: "{x. ~((j \<circ> f \<circ> h) x = x \<and> (j \<circ> g \<circ> h) x = x)} = j ` {x. (~ (f x = x \<and> g x = x))}"
+ have *: "{x. \<not>((j \<circ> f \<circ> h) x = x \<and> (j \<circ> g \<circ> h) x = x)} = j ` {x. (\<not> (f x = x \<and> g x = x))}"
using hj by (auto simp: jf jg image_iff, metis+)
ultimately show "bounded {x. \<not> ((j \<circ> f \<circ> h) x = x \<and> (j \<circ> g \<circ> 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 \<noteq> {}" "finite K" "K \<subseteq> S" and S: "S \<subseteq> T" "T \<subseteq> affine hull S" "connected S"
- obtains f g where "homeomorphism T T f g" "{x. (~ (f x = x \<and> g x = x))} \<subseteq> S"
- "bounded {x. (~ (f x = x \<and> g x = x))}" "\<And>x. x \<in> K \<Longrightarrow> f x \<in> U"
+ obtains f g where "homeomorphism T T f g" "{x. (\<not> (f x = x \<and> g x = x))} \<subseteq> S"
+ "bounded {x. (\<not> (f x = x \<and> g x = x))}" "\<And>x. x \<in> K \<Longrightarrow> f x \<in> U"
proof (cases "2 \<le> aff_dim S")
case True
have opeU': "openin (subtopology euclidean (affine hull S)) U"
@@ -8208,7 +8208,7 @@
then obtain \<gamma> where \<gamma>: "bij_betw \<gamma> K P"
using \<open>finite K\<close> finite_same_card_bij by blast
have "\<exists>f g. homeomorphism T T f g \<and> (\<forall>i \<in> K. f(id i) = \<gamma> i) \<and>
- {x. ~ (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. ~ (f x = x \<and> g x = x)}"
+ {x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. \<not> (f x = x \<and> g x = x)}"
proof (rule homeomorphism_moving_points_exists_gen [OF \<open>finite K\<close> _ _ True opeS S])
show "\<And>i. i \<in> K \<Longrightarrow> id i \<in> S \<and> \<gamma> i \<in> S"
by (metis id_apply opeU openin_contains_cball subsetCE \<open>P \<subseteq> U\<close> \<open>bij_betw \<gamma> K P\<close> \<open>K \<subseteq> S\<close> bij_betwE)
@@ -8314,10 +8314,10 @@
next
have "compact (j ` closure {x. \<not> (f x = x \<and> g x = x)})"
using bou by (auto simp: compact_continuous_image cont_hj)
- then have "bounded (j ` {x. (~ (f x = x \<and> g x = x))})"
+ then have "bounded (j ` {x. \<not> (f x = x \<and> g x = x)})"
by (rule bounded_closure_image [OF compact_imp_bounded])
moreover
- have *: "{x \<in> affine hull S. j (f (h x)) \<noteq> x \<or> j (g (h x)) \<noteq> x} = j ` {x. (~ (f x = x \<and> g x = x))}"
+ have *: "{x \<in> affine hull S. j (f (h x)) \<noteq> x \<or> j (g (h x)) \<noteq> x} = j ` {x. (\<not> (f x = x \<and> g x = x))}"
using h j by (auto simp: image_iff; metis)
ultimately have "bounded {x \<in> affine hull S. j (f (h x)) \<noteq> x \<or> j (g (h x)) \<noteq> x}"
by metis
@@ -8337,7 +8337,7 @@
qed
qed
-subsection\<open>nullhomotopic mappings\<close>
+subsection\<open>Nullhomotopic mappings\<close>
text\<open> 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 \<open>\<le>\<close> 0, and
--- 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: "~ (\<exists>k\<le>n. b k \<noteq> 0) \<Longrightarrow> \<exists>k. k \<noteq> 0 \<and> k \<le> Suc n \<and> c k \<noteq> 0"
+ then have extr_prem: "\<not> (\<exists>k\<le>n. b k \<noteq> 0) \<Longrightarrow> \<exists>k. k \<noteq> 0 \<and> k \<le> Suc n \<and> c k \<noteq> 0"
by (metis Suc.prems le0 minus_zero mult_zero_right)
have "\<exists>k\<le>n. b k \<noteq> 0"
apply (rule ccontr)
@@ -205,7 +205,7 @@
shows "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = 0) \<longleftrightarrow> (\<forall>k. k \<le> n \<longrightarrow> c k = 0)"
proof (cases "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = 0)")
case True
- then have "~ finite {z. (\<Sum>i\<le>n. c i * z^i) = 0}"
+ then have "\<not> finite {z. (\<Sum>i\<le>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)
--- 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 \<subseteq> S" and dis: "~disjnt (affine hull T) (rel_interior U)"
+ assumes T: "T face_of S" and "convex S" "U \<subseteq> S" and dis: "\<not> disjnt (affine hull T) (rel_interior U)"
shows "U \<subseteq> T"
apply (rule subset_of_face_of [OF T \<open>U \<subseteq> S\<close>])
using face_of_imp_eq_affine_Int [OF \<open>convex S\<close> 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 \<subseteq> - (affine hull S \<inter> {x. P x}) \<union> affine hull S \<inter> {x. Q x} \<Longrightarrow> S \<subseteq> {x. ~P x \<or> Q x}"
+ have *: "S \<subseteq> - (affine hull S \<inter> {x. P x}) \<union> affine hull S \<inter> {x. Q x} \<Longrightarrow> S \<subseteq> {x. \<not> P x \<or> Q x}"
for P Q
using hull_subset by fastforce
- have "S \<subseteq> {x. ~ (a' \<bullet> x \<le> b') \<or> a' \<bullet> x = b'}"
+ have "S \<subseteq> {x. \<not> (a' \<bullet> x \<le> b') \<or> a' \<bullet> 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)} \<subseteq> 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]: "\<not> (x extreme_point_of {})"
by (simp add: extreme_point_of_def)
lemma%unimportant extreme_point_of_singleton [iff]: "x extreme_point_of {a} \<longleftrightarrow> x = a"
@@ -974,10 +974,10 @@
(infixr "(facet'_of)" 50)
where "F facet_of S \<longleftrightarrow> F face_of S \<and> F \<noteq> {} \<and> aff_dim F = aff_dim S - 1"
-lemma%unimportant facet_of_empty [simp]: "~ S facet_of {}"
+lemma%unimportant facet_of_empty [simp]: "\<not> 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]: "\<not> S facet_of S "
by (simp add: facet_of_def)
lemma%unimportant facet_of_imp_face_of: "F facet_of S \<Longrightarrow> 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) \<and> ~ (norm b < norm x)"
+ ultimately have "\<not> (norm a < norm x) \<and> \<not> (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
+ "\<not> affine_dependent S
\<Longrightarrow> (x extreme_point_of (convex hull S) \<longleftrightarrow> x \<in> 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 "\<not> affine_dependent S"
shows "(T face_of (convex hull S) \<longleftrightarrow> (\<exists>c. c \<subseteq> S \<and> 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 "\<not> affine_dependent S"
shows "T facet_of (convex hull S) \<longleftrightarrow>
T \<noteq> {} \<and> (\<exists>u. u \<in> S \<and> 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 "\<not> affine_dependent c"
using \<open>c \<subseteq> S\<close> 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
+ "\<not>affine_dependent S
\<Longrightarrow> (T facet_of (convex hull S) \<longleftrightarrow>
2 \<le> card S \<and> (\<exists>u. u \<in> S \<and> T = convex hull (S - {u})))"
apply (simp add: facet_of_convex_hull_affine_independent)
@@ -1897,20 +1897,20 @@
have "\<exists>a' b'. a' \<noteq> 0 \<and>
affine hull S \<inter> {x. a' \<bullet> x \<le> b'} = affine hull S \<inter> h \<and>
(\<forall>w \<in> affine hull S. (w + a') \<in> affine hull S)"
- if "h \<in> F" "~(affine hull S \<subseteq> h)" for h
+ if "h \<in> F" "\<not>(affine hull S \<subseteq> h)" for h
proof -
have "a h \<noteq> 0" and "h = {x. a h \<bullet> x \<le> b h}" "h \<inter> \<Inter>F = \<Inter>F"
using \<open>h \<in> F\<close> ab by auto
then have "(affine hull S) \<inter> {x. a h \<bullet> x \<le> b h} \<noteq> {}"
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 \<subseteq> {x. a h \<bullet> x \<le> b h})"
+ moreover have "\<not> (affine hull S \<subseteq> {x. a h \<bullet> x \<le> b h})"
using \<open>h = {x. a h \<bullet> x \<le> b h}\<close> that(2) by blast
ultimately show ?thesis
using affine_parallel_slice [of "affine hull S"]
by (metis \<open>h = {x. a h \<bullet> x \<le> b h}\<close> affine_affine_hull)
qed
then obtain a b
- where ab: "\<And>h. \<lbrakk>h \<in> F; ~ (affine hull S \<subseteq> h)\<rbrakk>
+ where ab: "\<And>h. \<lbrakk>h \<in> F; \<not> (affine hull S \<subseteq> h)\<rbrakk>
\<Longrightarrow> a h \<noteq> 0 \<and>
affine hull S \<inter> {x. a h \<bullet> x \<le> b h} = affine hull S \<inter> h \<and>
(\<forall>w \<in> affine hull S. (w + a h) \<in> affine hull S)"
@@ -1918,7 +1918,7 @@
have seq2: "S = affine hull S \<inter> (\<Inter>h\<in>{h \<in> F. \<not> affine hull S \<subseteq> h}. {x. a h \<bullet> x \<le> b h})"
by (subst seq) (auto simp: ab INT_extend_simps)
show ?thesis
- apply (rule_tac x="(\<lambda>h. {x. a h \<bullet> x \<le> b h}) ` {h. h \<in> F \<and> ~(affine hull S \<subseteq> h)}" in exI)
+ apply (rule_tac x="(\<lambda>h. {x. a h \<bullet> x \<le> b h}) ` {h. h \<in> F \<and> \<not>(affine hull S \<subseteq> h)}" in exI)
apply (intro conjI seq2)
using \<open>finite F\<close> 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 \<and> ?P g \<and> ?Q g)" if "card g < n" for g
+ then have "\<not> (finite g \<and> ?P g \<and> ?Q g)" if "card g < n" for g
using that by (auto simp: n_def dest!: not_less_Least)
- then have *: "~ (?P g \<and> ?Q g)" if "g \<subset> F" for g
+ then have *: "\<not> (?P g \<and> ?Q g)" if "g \<subset> F" for g
using that \<open>finite F\<close> psubset_card_mono \<open>card F = n\<close>
by (metis finite_Int inf.strict_order_iff)
have 1: "\<And>F'. F' \<subset> F \<Longrightarrow> S \<subseteq> affine hull S \<inter> \<Inter>F'"
@@ -2230,7 +2230,7 @@
using seq by blast
have "F \<noteq> {}" using assms
by (metis affine_Int affine_Inter affine_affine_hull ex_in_conv face_of_affine_trivial)
- then obtain i where "i \<in> F" "~ (a i \<bullet> x < b i)"
+ then obtain i where "i \<in> F" "\<not> (a i \<bullet> x < b i)"
using \<open>x \<in> S\<close> rels xnot by auto
with xint have "a i \<bullet> 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 "\<not> affine_dependent S"
shows "(convex hull S) - rel_interior(convex hull S) = (\<Union>a\<in>S. convex hull (S - {a}))"
proof -
have "finite S" by (metis assms aff_independent_finite)
@@ -3074,17 +3074,17 @@
text\<open>The notion of n-simplex for integer @{term"n \<ge> -1"}\<close>
definition simplex :: "int \<Rightarrow> 'a::euclidean_space set \<Rightarrow> bool" (infix "simplex" 50)
- where "n simplex S \<equiv> \<exists>C. ~(affine_dependent C) \<and> int(card C) = n + 1 \<and> S = convex hull C"
+ where "n simplex S \<equiv> \<exists>C. \<not> affine_dependent C \<and> int(card C) = n + 1 \<and> S = convex hull C"
lemma simplex:
"n simplex S \<longleftrightarrow> (\<exists>C. finite C \<and>
- ~(affine_dependent C) \<and>
+ \<not> affine_dependent C \<and>
int(card C) = n + 1 \<and>
S = convex hull C)"
by (auto simp add: simplex_def intro: aff_independent_finite)
lemma simplex_convex_hull:
- "~affine_dependent C \<and> int(card C) = n + 1 \<Longrightarrow> n simplex (convex hull C)"
+ "\<not> affine_dependent C \<and> int(card C) = n + 1 \<Longrightarrow> n simplex (convex hull C)"
by (auto simp add: simplex_def)
lemma convex_simplex: "n simplex S \<Longrightarrow> convex S"
@@ -3170,7 +3170,7 @@
assumes "n simplex S" and a: "a \<notin> 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" "\<not> 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 \<noteq> z"
+ have "\<not>(between (t,u) z | between (u,z) t | between (z,t) u)" if "x \<noteq> 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 \<noteq> z"
+ then have "\<not> collinear {t, z, u}" if "x \<noteq> 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: \<N>_def poly\<M> polytope_imp_convex polytope_imp_closed)
have in_rel_interior: "(SOME z. z \<in> rel_interior C) \<in> rel_interior C" if "C \<in> \<N>" for C
using that poly\<M> polytope_imp_convex rel_interior_aff_dim some_in_eq by (fastforce simp: \<N>_def)
- have *: "\<exists>T. ~affine_dependent T \<and> card T \<le> n \<and> aff_dim K < n \<and> K = convex hull T"
+ have *: "\<exists>T. \<not> affine_dependent T \<and> card T \<le> n \<and> aff_dim K < n \<and> K = convex hull T"
if "K \<in> \<U>" for K
proof -
obtain r where r: "r simplex K"
--- 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 \<forall>C \<in> components(frontier S). ~bounded C"
+ else \<forall>C \<in> components(frontier S). \<not> bounded C"
proof -
have "S = {} \<or> 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" "\<And>C. C \<in> components (- S) \<Longrightarrow> ~bounded C"
+ assumes "connected S" "\<And>C. C \<in> components (- S) \<Longrightarrow> \<not> bounded C"
shows "inside S = {}"
using assms by (auto simp: components_def inside_def)
--- 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 \<subseteq> s"
+ assumes "\<not> affine_dependent s" "t \<subseteq> s"
shows "convex hull t = affine hull t \<inter> 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 "\<not> 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: "\<not> 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 "\<not> affine_dependent s"
shows "rel_interior(convex hull s) =
{y. \<exists>u. (\<forall>x \<in> s. 0 < u x) \<and> sum u s = 1 \<and> sum (\<lambda>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
+ "\<not> affine_dependent s
==> interior(convex hull s) =
(if card(s) \<le> DIM('a) then {}
else {y. \<exists>u. (\<forall>x \<in> s. 0 < u x) \<and> sum u s = 1 \<and> (\<Sum>x\<in>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 "\<not> affine_dependent s"
shows
"interior(convex hull s) =
(if card(s) \<le> DIM('a) then {}
@@ -4082,7 +4082,7 @@
lemma rel_frontier_convex_hull_explicit:
fixes s :: "'a::euclidean_space set"
- assumes "~ affine_dependent s"
+ assumes "\<not> affine_dependent s"
shows "rel_frontier(convex hull s) =
{y. \<exists>u. (\<forall>x \<in> s. 0 \<le> u x) \<and> (\<exists>x \<in> s. u x = 0) \<and> sum u s = 1 \<and> sum (\<lambda>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 "\<not> affine_dependent s"
shows "frontier(convex hull s) =
{y. \<exists>u. (\<forall>x \<in> s. 0 \<le> u x) \<and> (DIM ('a) < card s \<longrightarrow> (\<exists>x \<in> s. u x = 0)) \<and>
sum u s = 1 \<and> sum (\<lambda>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 "\<not> affine_dependent s"
shows "rel_frontier(convex hull s) = \<Union>{convex hull (s - {x}) |x. x \<in> 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 "\<not> affine_dependent s"
shows "frontier(convex hull s) =
(if card s \<le> 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 "\<not> affine_dependent s"
shows "frontier(convex hull s) =
(if card s \<le> DIM ('a) then convex hull s else \<Union>{convex hull (s - {x}) |x. x \<in> 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 "\<lbrakk>~ collinear{a,b,c}; DIM('a) = 2\<rbrakk>
+ shows "\<lbrakk>\<not> collinear{a,b,c}; DIM('a) = 2\<rbrakk>
\<Longrightarrow> interior(convex hull {a,b,c}) =
{v. \<exists>x y z. 0 < x \<and> 0 < y \<and> 0 < z \<and> x + y + z = 1 \<and>
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 "\<not>(affine_dependent S)"
shows "affine_dependent(insert a S) \<longleftrightarrow> a \<notin> S \<and> a \<in> affine hull S"
(is "?lhs = ?rhs")
proof safe
@@ -5887,7 +5887,7 @@
lemma affine_independent_insert:
fixes a :: "'a :: euclidean_space"
- shows "\<lbrakk>~(affine_dependent S); a \<notin> affine hull S\<rbrakk> \<Longrightarrow> ~(affine_dependent(insert a S))"
+ shows "\<lbrakk>\<not> affine_dependent S; a \<notin> affine hull S\<rbrakk> \<Longrightarrow> \<not> 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 \<union> t))"
+ assumes "\<not> affine_dependent(s \<union> t)"
shows convex_hull_Int_subset: "convex hull s \<inter> convex hull t \<subseteq> convex hull (s \<inter> t)" (is ?C)
and affine_hull_Int_subset: "affine hull s \<inter> affine hull t \<subseteq> affine hull (s \<inter> t)" (is ?A)
proof -
@@ -6201,7 +6201,7 @@
proposition affine_hull_Int:
fixes s :: "'a::euclidean_space set"
- assumes "~ (affine_dependent(s \<union> t))"
+ assumes "\<not> affine_dependent(s \<union> t)"
shows "affine hull (s \<inter> t) = affine hull s \<inter> 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 \<union> t))"
+ assumes "\<not> affine_dependent(s \<union> t)"
shows "convex hull (s \<inter> t) = convex hull s \<inter> 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 (\<Union>s))"
+ assumes "\<not> affine_dependent (\<Union>s)"
shows affine_hull_Inter: "affine hull (\<Inter>s) = (\<Inter>t\<in>s. affine hull t)" (is "?A")
and convex_hull_Inter: "convex hull (\<Inter>s) = (\<Inter>t\<in>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 \<in> convex hull S"
+ assumes naff: "\<not> affine_dependent S" and a: "a \<in> convex hull S"
and S: "T \<subseteq> S" "T' \<subseteq> S"
and x: "x \<in> convex hull (insert a T)"
and x': "x \<in> convex hull (insert a T')"
@@ -6395,7 +6395,7 @@
corollary convex_hull_exchange_Int:
fixes a :: "'a::euclidean_space"
- assumes "~ affine_dependent S" "a \<in> convex hull S" "T \<subseteq> S" "T' \<subseteq> S"
+ assumes "\<not> affine_dependent S" "a \<in> convex hull S" "T \<subseteq> S" "T' \<subseteq> S"
shows "(convex hull (insert a T)) \<inter> (convex hull (insert a T')) =
convex hull (insert a (T \<inter> T'))"
apply (rule subset_antisym)
@@ -6404,7 +6404,7 @@
lemma Int_closed_segment:
fixes b :: "'a::euclidean_space"
- assumes "b \<in> closed_segment a c \<or> ~collinear{a,b,c}"
+ assumes "b \<in> closed_segment a c \<or> \<not> collinear{a,b,c}"
shows "closed_segment a b \<inter> 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 \<inter> {x. a \<bullet> x \<le> b} \<noteq> {}"
- and "~ (S \<subseteq> {x. a \<bullet> x \<le> b})"
+ and "\<not> (S \<subseteq> {x. a \<bullet> x \<le> b})"
obtains a' b' where "a' \<noteq> 0"
"S \<inter> {x. a' \<bullet> x \<le> b'} = S \<inter> {x. a \<bullet> x \<le> b}"
"S \<inter> {x. a' \<bullet> x = b'} = S \<inter> {x. a \<bullet> x = b}"
@@ -7669,7 +7669,7 @@
assumes opS: "openin (subtopology euclidean (S \<union> T)) S"
and opT: "openin (subtopology euclidean (S \<union> T)) T"
and contf: "continuous_on S f" and contg: "continuous_on T g"
- and fg: "\<And>x. x \<in> S \<and> ~P x \<or> x \<in> T \<and> P x \<Longrightarrow> f x = g x"
+ and fg: "\<And>x. x \<in> S \<and> \<not>P x \<or> x \<in> T \<and> P x \<Longrightarrow> f x = g x"
shows "continuous_on (S \<union> T) (\<lambda>x. if P x then f x else g x)"
proof -
have "\<And>x. x \<in> S \<Longrightarrow> (if P x then f x else g x) = f x" "\<And>x. x \<in> T \<Longrightarrow> (if P x then f x else g x) = g x"
--- 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 \<in> Basis"
shows "division_points (cbox a b \<inter> {x. x\<bullet>k \<le> c}) {l \<inter> {x. x\<bullet>k \<le> c} | l . l \<in> d \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}} \<subseteq>
division_points (cbox a b) d" (is ?t1)
- and "division_points (cbox a b \<inter> {x. x\<bullet>k \<ge> c}) {l \<inter> {x. x\<bullet>k \<ge> c} | l . l \<in> d \<and> ~(l \<inter> {x. x\<bullet>k \<ge> c} = {})} \<subseteq>
+ and "division_points (cbox a b \<inter> {x. x\<bullet>k \<ge> c}) {l \<inter> {x. x\<bullet>k \<ge> c} | l . l \<in> d \<and> \<not>(l \<inter> {x. x\<bullet>k \<ge> c} = {})} \<subseteq>
division_points (cbox a b) d" (is ?t2)
proof%unimportant -
note assm = division_ofD[OF assms(1)]
@@ -2536,7 +2536,7 @@
show "\<And>K. K \<in> ?D1 \<Longrightarrow> finite {L. L \<in> ?D1 \<and> K \<subseteq> L}"
using fin by simp (metis (mono_tags, lifting) Collect_mono rev_finite_subset)
qed
- let ?\<D> = "{K \<in> \<D>. \<forall>K'. K' \<in> \<D> \<and> K \<noteq> K' \<longrightarrow> ~(K \<subseteq> K')}"
+ let ?\<D> = "{K \<in> \<D>. \<forall>K'. K' \<in> \<D> \<and> K \<noteq> K' \<longrightarrow> \<not>(K \<subseteq> K')}"
show ?thesis
proof (rule that)
show "countable ?\<D>"
--- 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 \<longleftrightarrow>
- ~(\<exists>E1 E2. openin (subtopology euclidean S) E1 \<and>
+ \<not>(\<exists>E1 E2. openin (subtopology euclidean S) E1 \<and>
openin (subtopology euclidean S) E2 \<and>
S \<subseteq> E1 \<union> E2 \<and> E1 \<inter> E2 = {} \<and> E1 \<noteq> {} \<and> E2 \<noteq> {})"
apply (simp add: connected_def openin_open disjoint_iff_not_equal, safe)
@@ -1113,7 +1113,7 @@
lemma connected_openin_eq:
"connected S \<longleftrightarrow>
- ~(\<exists>E1 E2. openin (subtopology euclidean S) E1 \<and>
+ \<not>(\<exists>E1 E2. openin (subtopology euclidean S) E1 \<and>
openin (subtopology euclidean S) E2 \<and>
E1 \<union> E2 = S \<and> E1 \<inter> E2 = {} \<and>
E1 \<noteq> {} \<and> E2 \<noteq> {})"
@@ -1152,7 +1152,7 @@
lemma connected_closedin_eq:
"connected S \<longleftrightarrow>
- ~(\<exists>E1 E2.
+ \<not>(\<exists>E1 E2.
closedin (subtopology euclidean S) E1 \<and>
closedin (subtopology euclidean S) E2 \<and>
E1 \<union> E2 = S \<and> E1 \<inter> E2 = {} \<and>
@@ -1946,7 +1946,7 @@
by (auto simp: box_def inner_sum_left inner_Basis sum.If_cases)
qed
-subsection \<open>Intervals in general, including infinite and mixtures of open and closed\<close>
+subsection \<open>General Intervals\<close>
definition%important "is_interval (s::('a::euclidean_space) set) \<longleftrightarrow>
(\<forall>a\<in>s. \<forall>b\<in>s. \<forall>x. (\<forall>i\<in>Basis. ((a\<bullet>i \<le> x\<bullet>i \<and> x\<bullet>i \<le> b\<bullet>i) \<or> (b\<bullet>i \<le> x\<bullet>i \<and> x\<bullet>i \<le> a\<bullet>i))) \<longrightarrow> x \<in> s)"
@@ -3645,7 +3645,7 @@
corollary cobounded_imp_unbounded:
fixes S :: "'a::{real_normed_vector, perfect_space} set"
- shows "bounded (- S) \<Longrightarrow> ~ (bounded S)"
+ shows "bounded (- S) \<Longrightarrow> \<not> 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 "\<not> compact (UNIV::'a set)"
by (simp add: compact_eq_bounded_closed)
text\<open>Representing sets as the union of a chain of compact sets.\<close>
--- 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" "\<And>c. P(\<lambda>x. c::real)"
"\<And>f. P f \<Longrightarrow> continuous_on S f"
"\<And>f g. P(f) \<and> P(g) \<Longrightarrow> P(\<lambda>x. f x + g x)" "\<And>f g. P(f) \<and> P(g) \<Longrightarrow> P(\<lambda>x. f x * g x)"
- "\<And>x y. x \<in> S \<and> y \<in> S \<and> ~(x = y) \<Longrightarrow> \<exists>f. P(f) \<and> ~(f x = f y)"
+ "\<And>x y. x \<in> S \<and> y \<in> S \<and> x \<noteq> y \<Longrightarrow> \<exists>f. P(f) \<and> f x \<noteq> f y"
"continuous_on S f"
"0 < e"
shows "\<exists>g. P(g) \<and> (\<forall>x \<in> S. \<bar>f x - g x\<bar> < e)"
--- 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 \<in> interior(convex hull {a,b,c})"
- shows "~ (Im(a/b) \<le> 0 \<and> 0 \<le> Im(b/c))"
+ shows "\<not> (Im(a/b) \<le> 0 \<and> 0 \<le> Im(b/c))"
proof -
{ assume 0: "Im(a/b) \<le> 0" "0 \<le> Im(b/c)"
have "0 \<notin> interior (convex hull {a,b,c})"
@@ -224,7 +224,7 @@
by (simp_all add: assms path_image_join)
with \<open>0 < e\<close> dac have "c \<notin> affine hull {a - of_real e, a + of_real e}"
by (simp add: segment_as_ball not_le)
- with \<open>0 < e\<close> have *: "~collinear{a - e, c,a + e}"
+ with \<open>0 < e\<close> have *: "\<not> 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) \<in> interior(convex hull {a - e, c, a + e})"