--- a/src/HOL/Analysis/Abstract_Topology.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Abstract_Topology.thy Fri Apr 12 22:09:25 2019 +0200
@@ -12,11 +12,11 @@
subsection \<open>General notion of a topology as a value\<close>
-definition%important istopology :: "('a set \<Rightarrow> bool) \<Rightarrow> bool" where
+definition\<^marker>\<open>tag important\<close> istopology :: "('a set \<Rightarrow> bool) \<Rightarrow> bool" where
"istopology L \<longleftrightarrow>
L {} \<and> (\<forall>S T. L S \<longrightarrow> L T \<longrightarrow> L (S \<inter> T)) \<and> (\<forall>K. Ball K L \<longrightarrow> L (\<Union>K))"
-typedef%important 'a topology = "{L::('a set) \<Rightarrow> bool. istopology L}"
+typedef\<^marker>\<open>tag important\<close> 'a topology = "{L::('a set) \<Rightarrow> bool. istopology L}"
morphisms "openin" "topology"
unfolding istopology_def by blast
@@ -116,7 +116,7 @@
subsubsection \<open>Closed sets\<close>
-definition%important closedin :: "'a topology \<Rightarrow> 'a set \<Rightarrow> bool" where
+definition\<^marker>\<open>tag important\<close> closedin :: "'a topology \<Rightarrow> 'a set \<Rightarrow> bool" where
"closedin U S \<longleftrightarrow> S \<subseteq> topspace U \<and> openin U (topspace U - S)"
lemma closedin_subset: "closedin U S \<Longrightarrow> S \<subseteq> topspace U"
@@ -240,7 +240,7 @@
subsection \<open>Subspace topology\<close>
-definition%important subtopology :: "'a topology \<Rightarrow> 'a set \<Rightarrow> 'a topology" where
+definition\<^marker>\<open>tag important\<close> subtopology :: "'a topology \<Rightarrow> 'a set \<Rightarrow> 'a topology" where
"subtopology U V = topology (\<lambda>T. \<exists>S. T = S \<inter> V \<and> openin U S)"
lemma istopology_subtopology: "istopology (\<lambda>T. \<exists>S. T = S \<inter> V \<and> openin U S)"
@@ -415,7 +415,7 @@
subsection \<open>The canonical topology from the underlying type class\<close>
-abbreviation%important euclidean :: "'a::topological_space topology"
+abbreviation\<^marker>\<open>tag important\<close> euclidean :: "'a::topological_space topology"
where "euclidean \<equiv> topology open"
abbreviation top_of_set :: "'a::topological_space set \<Rightarrow> 'a topology"
@@ -1416,7 +1416,7 @@
by (simp add: Sup_le_iff closure_of_minimal)
-subsection%important \<open>Continuous maps\<close>
+subsection\<^marker>\<open>tag important\<close> \<open>Continuous maps\<close>
text \<open>We will need to deal with continuous maps in terms of topologies and not in terms
of type classes, as defined below.\<close>
@@ -1728,7 +1728,7 @@
declare continuous_map_id_subt [unfolded id_def, simp]
-lemma%important continuous_map_alt:
+lemma\<^marker>\<open>tag important\<close> continuous_map_alt:
"continuous_map T1 T2 f
= ((\<forall>U. openin T2 U \<longrightarrow> openin T1 (f -` U \<inter> topspace T1)) \<and> f ` topspace T1 \<subseteq> topspace T2)"
by (auto simp: continuous_map_def vimage_def image_def Collect_conj_eq inf_commute)
@@ -3793,7 +3793,7 @@
shows "continuous_map euclidean (top_of_set S) f"
by (simp add: cont continuous_map_into_subtopology eq image_subset_iff_subset_vimage)
-subsection%unimportant \<open>Half-global and completely global cases\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Half-global and completely global cases\<close>
lemma continuous_openin_preimage_gen:
assumes "continuous_on S f" "open T"
@@ -4027,7 +4027,7 @@
using assms by (simp add: Lim_transform_within_openin continuous_within)
-subsection%important \<open>The topology generated by some (open) subsets\<close>
+subsection\<^marker>\<open>tag important\<close> \<open>The topology generated by some (open) subsets\<close>
text \<open>In the definition below of a generated topology, the \<open>Empty\<close> case is not necessary,
as it follows from \<open>UN\<close> taking for \<open>K\<close> the empty set. However, it is convenient to have,
@@ -4057,7 +4057,7 @@
using assms(3) apply (induct rule: generate_topology_on.induct)
using assms(1) assms(2) unfolding istopology_def by auto
-abbreviation%unimportant topology_generated_by::"('a set set) \<Rightarrow> ('a topology)"
+abbreviation\<^marker>\<open>tag unimportant\<close> topology_generated_by::"('a set set) \<Rightarrow> ('a topology)"
where "topology_generated_by S \<equiv> topology (generate_topology_on S)"
lemma openin_topology_generated_by_iff:
@@ -4254,7 +4254,7 @@
using assms continuous_on_generated_topo_iff by blast
-subsection%important \<open>Pullback topology\<close>
+subsection\<^marker>\<open>tag important\<close> \<open>Pullback topology\<close>
text \<open>Pulling back a topology by map gives again a topology. \<open>subtopology\<close> is
a special case of this notion, pulling back by the identity. We introduce the general notion as
@@ -4264,7 +4264,7 @@
text \<open>\<open>pullback_topology A f T\<close> is the pullback of the topology \<open>T\<close> by the map \<open>f\<close> on
the set \<open>A\<close>.\<close>
-definition%important pullback_topology::"('a set) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('b topology) \<Rightarrow> ('a topology)"
+definition\<^marker>\<open>tag important\<close> pullback_topology::"('a set) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('b topology) \<Rightarrow> ('a topology)"
where "pullback_topology A f T = topology (\<lambda>S. \<exists>U. openin T U \<and> S = f-`U \<inter> A)"
lemma istopology_pullback_topology:
--- a/src/HOL/Analysis/Abstract_Topology_2.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Abstract_Topology_2.thy Fri Apr 12 22:09:25 2019 +0200
@@ -281,7 +281,7 @@
with \<open>x = f y\<close> show "x \<in> f ` interior S" ..
qed
-subsection%unimportant \<open>Equality of continuous functions on closure and related results\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Equality of continuous functions on closure and related results\<close>
lemma continuous_closedin_preimage_constant:
fixes f :: "_ \<Rightarrow> 'b::t1_space"
@@ -319,7 +319,7 @@
then show ?thesis by auto
qed
-subsection%unimportant \<open>A function constant on a set\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>A function constant on a set\<close>
definition constant_on (infixl "(constant'_on)" 50)
where "f constant_on A \<equiv> \<exists>y. \<forall>x\<in>A. f x = y"
@@ -341,7 +341,7 @@
by metis
-subsection%unimportant \<open>Continuity relative to a union.\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Continuity relative to a union.\<close>
lemma continuous_on_Un_local:
"\<lbrakk>closedin (top_of_set (s \<union> t)) s; closedin (top_of_set (s \<union> t)) t;
@@ -391,7 +391,7 @@
by (auto simp: continuous_on_id intro: continuous_on_cases_le [where h = id, simplified])
-subsection%unimportant\<open>Inverse function property for open/closed maps\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Inverse function property for open/closed maps\<close>
lemma continuous_on_inverse_open_map:
assumes contf: "continuous_on S f"
@@ -479,7 +479,7 @@
by (simp add: continuous_on_closed oo)
qed
-subsection%unimportant \<open>Seperability\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Seperability\<close>
lemma subset_second_countable:
obtains \<B> :: "'a:: second_countable_topology set set"
@@ -540,7 +540,7 @@
qed
-subsection%unimportant\<open>Closed Maps\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Closed Maps\<close>
lemma continuous_imp_closed_map:
fixes f :: "'a::t2_space \<Rightarrow> 'b::t2_space"
@@ -561,7 +561,7 @@
by (fastforce simp add: closedin_closed)
qed
-subsection%unimportant\<open>Open Maps\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Open Maps\<close>
lemma open_map_restrict:
assumes opeU: "openin (top_of_set (S \<inter> f -` T')) U"
@@ -576,7 +576,7 @@
qed
-subsection%unimportant\<open>Quotient maps\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Quotient maps\<close>
lemma quotient_map_imp_continuous_open:
assumes T: "f ` S \<subseteq> T"
@@ -693,7 +693,7 @@
openin (top_of_set T) U"
by (metis (no_types, lifting) assms closed_map_imp_quotient_map continuous_imp_closed_map)
-subsection%unimportant\<open>Pasting lemmas for functions, for of casewise definitions\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Pasting lemmas for functions, for of casewise definitions\<close>
subsubsection\<open>on open sets\<close>
@@ -905,11 +905,11 @@
subsection \<open>Retractions\<close>
-definition%important retraction :: "('a::topological_space) set \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> retraction :: "('a::topological_space) set \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool"
where "retraction S T r \<longleftrightarrow>
T \<subseteq> S \<and> continuous_on S r \<and> r ` S \<subseteq> T \<and> (\<forall>x\<in>T. r x = x)"
-definition%important retract_of (infixl "retract'_of" 50) where
+definition\<^marker>\<open>tag important\<close> retract_of (infixl "retract'_of" 50) where
"T retract_of S \<longleftrightarrow> (\<exists>r. retraction S T r)"
lemma retraction_idempotent: "retraction S T r \<Longrightarrow> x \<in> S \<Longrightarrow> r (r x) = r x"
--- a/src/HOL/Analysis/Arcwise_Connected.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Arcwise_Connected.thy Fri Apr 12 22:09:25 2019 +0200
@@ -106,7 +106,7 @@
qed (meson assms compact_imp_closed seq_compact_closed_subset seq_compact_eq_compact)+
-subsection%unimportant\<open>Arcwise Connections\<close>(*FIX ME this subsection is empty(?) *)
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Arcwise Connections\<close>(*FIX ME this subsection is empty(?) *)
subsection\<open>Density of points with dyadic rational coordinates\<close>
@@ -201,7 +201,7 @@
qed
-definition%unimportant dyadics :: "'a::field_char_0 set" where "dyadics \<equiv> \<Union>k m. {of_nat m / 2^k}"
+definition\<^marker>\<open>tag unimportant\<close> dyadics :: "'a::field_char_0 set" where "dyadics \<equiv> \<Union>k m. {of_nat m / 2^k}"
lemma real_in_dyadics [simp]: "real m \<in> dyadics"
apply (simp add: dyadics_def)
@@ -344,7 +344,7 @@
qed
-function%unimportant (domintros) dyad_rec :: "[nat \<Rightarrow> 'a, 'a\<Rightarrow>'a, 'a\<Rightarrow>'a, real] \<Rightarrow> 'a" where
+function\<^marker>\<open>tag unimportant\<close> (domintros) dyad_rec :: "[nat \<Rightarrow> 'a, 'a\<Rightarrow>'a, 'a\<Rightarrow>'a, real] \<Rightarrow> 'a" where
"dyad_rec b l r (real m) = b m"
| "dyad_rec b l r ((4 * real m + 1) / 2 ^ (Suc n)) = l (dyad_rec b l r ((2*m + 1) / 2^n))"
| "dyad_rec b l r ((4 * real m + 3) / 2 ^ (Suc n)) = r (dyad_rec b l r ((2*m + 1) / 2^n))"
@@ -479,12 +479,12 @@
finally show ?thesis .
qed
-definition%unimportant dyad_rec2
+definition\<^marker>\<open>tag unimportant\<close> dyad_rec2
where "dyad_rec2 u v lc rc x =
dyad_rec (\<lambda>z. (u,v)) (\<lambda>(a,b). (a, lc a b (midpoint a b))) (\<lambda>(a,b). (rc a b (midpoint a b), b)) (2*x)"
-abbreviation%unimportant leftrec where "leftrec u v lc rc x \<equiv> fst (dyad_rec2 u v lc rc x)"
-abbreviation%unimportant rightrec where "rightrec u v lc rc x \<equiv> snd (dyad_rec2 u v lc rc x)"
+abbreviation\<^marker>\<open>tag unimportant\<close> leftrec where "leftrec u v lc rc x \<equiv> fst (dyad_rec2 u v lc rc x)"
+abbreviation\<^marker>\<open>tag unimportant\<close> rightrec where "rightrec u v lc rc x \<equiv> snd (dyad_rec2 u v lc rc x)"
lemma leftrec_base: "leftrec u v lc rc (real m / 2) = u"
by (simp add: dyad_rec2_def)
--- a/src/HOL/Analysis/Ball_Volume.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Ball_Volume.thy Fri Apr 12 22:09:25 2019 +0200
@@ -14,7 +14,7 @@
dimension need not be an integer; we also allow fractional dimensions, although we do
not use this case or prove anything about it for now.
\<close>
-definition%important unit_ball_vol :: "real \<Rightarrow> real" where
+definition\<^marker>\<open>tag important\<close> unit_ball_vol :: "real \<Rightarrow> real" where
"unit_ball_vol n = pi powr (n / 2) / Gamma (n / 2 + 1)"
lemma unit_ball_vol_pos [simp]: "n \<ge> 0 \<Longrightarrow> unit_ball_vol n > 0"
@@ -177,7 +177,7 @@
assumes r: "r \<ge> 0"
begin
-theorem%unimportant emeasure_cball:
+theorem\<^marker>\<open>tag unimportant\<close> emeasure_cball:
"emeasure lborel (cball c r) = ennreal (unit_ball_vol (DIM('a)) * r ^ DIM('a))"
proof (cases "r = 0")
case False
@@ -202,11 +202,11 @@
finally show ?thesis .
qed auto
-corollary%unimportant content_cball:
+corollary\<^marker>\<open>tag unimportant\<close> content_cball:
"content (cball c r) = unit_ball_vol (DIM('a)) * r ^ DIM('a)"
by (simp add: measure_def emeasure_cball r)
-corollary%unimportant emeasure_ball:
+corollary\<^marker>\<open>tag unimportant\<close> emeasure_ball:
"emeasure lborel (ball c r) = ennreal (unit_ball_vol (DIM('a)) * r ^ DIM('a))"
proof -
from negligible_sphere[of c r] have "sphere c r \<in> null_sets lborel"
@@ -219,7 +219,7 @@
finally show ?thesis ..
qed
-corollary%important content_ball:
+corollary\<^marker>\<open>tag important\<close> content_ball:
"content (ball c r) = unit_ball_vol (DIM('a)) * r ^ DIM('a)"
by (simp add: measure_def r emeasure_ball)
@@ -289,25 +289,25 @@
lemma unit_ball_vol_1 [simp]: "unit_ball_vol 1 = 2"
using unit_ball_vol_odd[of 0] by simp
-corollary%unimportant
+corollary\<^marker>\<open>tag unimportant\<close>
unit_ball_vol_2: "unit_ball_vol 2 = pi"
and unit_ball_vol_3: "unit_ball_vol 3 = 4 / 3 * pi"
and unit_ball_vol_4: "unit_ball_vol 4 = pi\<^sup>2 / 2"
and unit_ball_vol_5: "unit_ball_vol 5 = 8 / 15 * pi\<^sup>2"
by (simp_all add: eval_unit_ball_vol)
-corollary%unimportant circle_area:
+corollary\<^marker>\<open>tag unimportant\<close> circle_area:
"r \<ge> 0 \<Longrightarrow> content (ball c r :: (real ^ 2) set) = r ^ 2 * pi"
by (simp add: content_ball unit_ball_vol_2)
-corollary%unimportant sphere_volume:
+corollary\<^marker>\<open>tag unimportant\<close> sphere_volume:
"r \<ge> 0 \<Longrightarrow> content (ball c r :: (real ^ 3) set) = 4 / 3 * r ^ 3 * pi"
by (simp add: content_ball unit_ball_vol_3)
text \<open>
Useful equivalent forms
\<close>
-corollary%unimportant content_ball_eq_0_iff [simp]: "content (ball c r) = 0 \<longleftrightarrow> r \<le> 0"
+corollary\<^marker>\<open>tag unimportant\<close> content_ball_eq_0_iff [simp]: "content (ball c r) = 0 \<longleftrightarrow> r \<le> 0"
proof -
have "r > 0 \<Longrightarrow> content (ball c r) > 0"
by (simp add: content_ball unit_ball_vol_def)
@@ -315,10 +315,10 @@
by (fastforce simp: ball_empty)
qed
-corollary%unimportant content_ball_gt_0_iff [simp]: "0 < content (ball z r) \<longleftrightarrow> 0 < r"
+corollary\<^marker>\<open>tag unimportant\<close> content_ball_gt_0_iff [simp]: "0 < content (ball z r) \<longleftrightarrow> 0 < r"
by (auto simp: zero_less_measure_iff)
-corollary%unimportant content_cball_eq_0_iff [simp]: "content (cball c r) = 0 \<longleftrightarrow> r \<le> 0"
+corollary\<^marker>\<open>tag unimportant\<close> content_cball_eq_0_iff [simp]: "content (cball c r) = 0 \<longleftrightarrow> r \<le> 0"
proof (cases "r = 0")
case False
moreover have "r > 0 \<Longrightarrow> content (cball c r) > 0"
@@ -327,7 +327,7 @@
by fastforce
qed auto
-corollary%unimportant content_cball_gt_0_iff [simp]: "0 < content (cball z r) \<longleftrightarrow> 0 < r"
+corollary\<^marker>\<open>tag unimportant\<close> content_cball_gt_0_iff [simp]: "0 < content (cball z r) \<longleftrightarrow> 0 < r"
by (auto simp: zero_less_measure_iff)
end
--- a/src/HOL/Analysis/Binary_Product_Measure.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Binary_Product_Measure.thy Fri Apr 12 22:09:25 2019 +0200
@@ -16,7 +16,7 @@
subsection "Binary products"
-definition%important pair_measure (infixr "\<Otimes>\<^sub>M" 80) where
+definition\<^marker>\<open>tag important\<close> pair_measure (infixr "\<Otimes>\<^sub>M" 80) where
"A \<Otimes>\<^sub>M B = measure_of (space A \<times> space B)
{a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}
(\<lambda>X. \<integral>\<^sup>+x. (\<integral>\<^sup>+y. indicator X (x,y) \<partial>B) \<partial>A)"
@@ -338,7 +338,7 @@
subsection \<open>Binary products of \<open>\<sigma>\<close>-finite emeasure spaces\<close>
-locale%unimportant pair_sigma_finite = M1?: sigma_finite_measure M1 + M2?: sigma_finite_measure M2
+locale\<^marker>\<open>tag unimportant\<close> pair_sigma_finite = M1?: sigma_finite_measure M1 + M2?: sigma_finite_measure M2
for M1 :: "'a measure" and M2 :: "'b measure"
lemma (in pair_sigma_finite) measurable_emeasure_Pair1:
@@ -394,7 +394,7 @@
qed
qed
-sublocale%unimportant pair_sigma_finite \<subseteq> P?: sigma_finite_measure "M1 \<Otimes>\<^sub>M M2"
+sublocale\<^marker>\<open>tag unimportant\<close> pair_sigma_finite \<subseteq> P?: sigma_finite_measure "M1 \<Otimes>\<^sub>M M2"
proof
from M1.sigma_finite_countable guess F1 ..
moreover from M2.sigma_finite_countable guess F2 ..
--- a/src/HOL/Analysis/Bochner_Integration.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Bochner_Integration.thy Fri Apr 12 22:09:25 2019 +0200
@@ -309,7 +309,7 @@
then show "emeasure M {y \<in> space M. f y \<noteq> 0} \<noteq> \<infinity>" by simp
qed fact
-definition%important simple_bochner_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> 'b::real_vector) \<Rightarrow> 'b" where
+definition\<^marker>\<open>tag important\<close> simple_bochner_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> 'b::real_vector) \<Rightarrow> 'b" where
"simple_bochner_integral M f = (\<Sum>y\<in>f`space M. measure M {x\<in>space M. f x = y} *\<^sub>R y)"
proposition simple_bochner_integral_partition:
@@ -895,7 +895,7 @@
end
-definition%important lebesgue_integral ("integral\<^sup>L") where
+definition\<^marker>\<open>tag important\<close> lebesgue_integral ("integral\<^sup>L") where
"integral\<^sup>L M f = (if \<exists>x. has_bochner_integral M f x then THE x. has_bochner_integral M f x else 0)"
syntax
--- a/src/HOL/Analysis/Borel_Space.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Borel_Space.thy Fri Apr 12 22:09:25 2019 +0200
@@ -25,7 +25,7 @@
by (auto intro: topological_basis_prod topological_basis_trivial topological_basis_imp_subbasis)
qed
-definition%important "mono_on f A \<equiv> \<forall>r s. r \<in> A \<and> s \<in> A \<and> r \<le> s \<longrightarrow> f r \<le> f s"
+definition\<^marker>\<open>tag important\<close> "mono_on f A \<equiv> \<forall>r s. r \<in> A \<and> s \<in> A \<and> r \<le> s \<longrightarrow> f r \<le> f s"
lemma mono_onI:
"(\<And>r s. r \<in> A \<Longrightarrow> s \<in> A \<Longrightarrow> r \<le> s \<Longrightarrow> f r \<le> f s) \<Longrightarrow> mono_on f A"
@@ -41,7 +41,7 @@
lemma mono_on_subset: "mono_on f A \<Longrightarrow> B \<subseteq> A \<Longrightarrow> mono_on f B"
unfolding mono_on_def by auto
-definition%important "strict_mono_on f A \<equiv> \<forall>r s. r \<in> A \<and> s \<in> A \<and> r < s \<longrightarrow> f r < f s"
+definition\<^marker>\<open>tag important\<close> "strict_mono_on f A \<equiv> \<forall>r s. r \<in> A \<and> s \<in> A \<and> r < s \<longrightarrow> f r < f s"
lemma strict_mono_onI:
"(\<And>r s. r \<in> A \<Longrightarrow> s \<in> A \<Longrightarrow> r < s \<Longrightarrow> f r < f s) \<Longrightarrow> strict_mono_on f A"
@@ -296,7 +296,7 @@
subsection \<open>Generic Borel spaces\<close>
-definition%important (in topological_space) borel :: "'a measure" where
+definition\<^marker>\<open>tag important\<close> (in topological_space) borel :: "'a measure" where
"borel = sigma UNIV {S. open S}"
abbreviation "borel_measurable M \<equiv> measurable M borel"
@@ -1508,7 +1508,7 @@
lemma borel_measurable_arctan [measurable]: "arctan \<in> borel_measurable borel"
by (intro borel_measurable_continuous_on1 continuous_intros)
-lemma%important borel_measurable_complex_iff:
+lemma\<^marker>\<open>tag important\<close> borel_measurable_complex_iff:
"f \<in> borel_measurable M \<longleftrightarrow>
(\<lambda>x. Re (f x)) \<in> borel_measurable M \<and> (\<lambda>x. Im (f x)) \<in> borel_measurable M"
apply auto
@@ -1701,7 +1701,7 @@
"f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. enn2real (f x)) \<in> M \<rightarrow>\<^sub>M borel"
unfolding enn2real_def[abs_def] by measurable
-definition%important [simp]: "is_borel f M \<longleftrightarrow> f \<in> borel_measurable M"
+definition\<^marker>\<open>tag important\<close> [simp]: "is_borel f M \<longleftrightarrow> f \<in> borel_measurable M"
lemma is_borel_transfer[transfer_rule]: "rel_fun (rel_fun (=) pcr_ennreal) (=) is_borel is_borel"
unfolding is_borel_def[abs_def]
--- a/src/HOL/Analysis/Bounded_Continuous_Function.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Bounded_Continuous_Function.thy Fri Apr 12 22:09:25 2019 +0200
@@ -7,7 +7,7 @@
subsection \<open>Definition\<close>
-definition%important "bcontfun = {f. continuous_on UNIV f \<and> bounded (range f)}"
+definition\<^marker>\<open>tag important\<close> "bcontfun = {f. continuous_on UNIV f \<and> bounded (range f)}"
typedef (overloaded) ('a, 'b) bcontfun ("(_ \<Rightarrow>\<^sub>C /_)" [22, 21] 21) =
"bcontfun::('a::topological_space \<Rightarrow> 'b::metric_space) set"
@@ -41,7 +41,7 @@
instantiation bcontfun :: (topological_space, metric_space) metric_space
begin
-lift_definition%important dist_bcontfun :: "'a \<Rightarrow>\<^sub>C 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>C 'b \<Rightarrow> real"
+lift_definition\<^marker>\<open>tag important\<close> dist_bcontfun :: "'a \<Rightarrow>\<^sub>C 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>C 'b \<Rightarrow> real"
is "\<lambda>f g. (SUP x. dist (f x) (g x))" .
definition uniformity_bcontfun :: "('a \<Rightarrow>\<^sub>C 'b \<times> 'a \<Rightarrow>\<^sub>C 'b) filter"
@@ -175,8 +175,8 @@
subsection \<open>Complete Space\<close>
-instance%important bcontfun :: (metric_space, complete_space) complete_space
-proof%unimportant
+instance\<^marker>\<open>tag important\<close> bcontfun :: (metric_space, complete_space) complete_space
+proof
fix f :: "nat \<Rightarrow> ('a, 'b) bcontfun"
assume "Cauchy f" \<comment> \<open>Cauchy equals uniform convergence\<close>
then obtain g where "uniform_limit UNIV f g sequentially"
@@ -191,9 +191,9 @@
qed
-subsection%unimportant \<open>Supremum norm for a normed vector space\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Supremum norm for a normed vector space\<close>
-instantiation%unimportant bcontfun :: (topological_space, real_normed_vector) real_vector
+instantiation\<^marker>\<open>tag unimportant\<close> bcontfun :: (topological_space, real_normed_vector) real_vector
begin
lemma uminus_cont: "f \<in> bcontfun \<Longrightarrow> (\<lambda>x. - f x) \<in> bcontfun" for f::"'a \<Rightarrow> 'b"
@@ -247,7 +247,7 @@
"bounded (range f) \<Longrightarrow> norm (f x) \<le> (SUP x. norm (f x))"
by (auto intro!: cSUP_upper bounded_imp_bdd_above simp: bounded_norm_comp)
-instantiation%unimportant bcontfun :: (topological_space, real_normed_vector) real_normed_vector
+instantiation\<^marker>\<open>tag unimportant\<close> bcontfun :: (topological_space, real_normed_vector) real_normed_vector
begin
definition norm_bcontfun :: "('a, 'b) bcontfun \<Rightarrow> real"
@@ -290,7 +290,7 @@
using dist_bound[of f 0 b] assms
by (simp add: dist_norm)
-subsection%unimportant \<open>(bounded) continuous extenstion\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>(bounded) continuous extenstion\<close>
lemma integral_clamp:
"integral {t0::real..clamp t0 t1 x} f =
--- a/src/HOL/Analysis/Bounded_Linear_Function.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Bounded_Linear_Function.thy Fri Apr 12 22:09:25 2019 +0200
@@ -39,7 +39,7 @@
lemmas onorm_componentwise_le = order_trans[OF onorm_componentwise]
-subsection%unimportant \<open>Intro rules for \<^term>\<open>bounded_linear\<close>\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Intro rules for \<^term>\<open>bounded_linear\<close>\<close>
named_theorems bounded_linear_intros
@@ -81,7 +81,7 @@
bounded_linear_inner_right_comp
-subsection%unimportant \<open>declaration of derivative/continuous/tendsto introduction rules for bounded linear functions\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>declaration of derivative/continuous/tendsto introduction rules for bounded linear functions\<close>
attribute_setup bounded_linear =
\<open>Scan.succeed (Thm.declaration_attribute (fn thm =>
@@ -116,7 +116,7 @@
subsection \<open>Type of bounded linear functions\<close>
-typedef%important (overloaded) ('a, 'b) blinfun ("(_ \<Rightarrow>\<^sub>L /_)" [22, 21] 21) =
+typedef\<^marker>\<open>tag important\<close> (overloaded) ('a, 'b) blinfun ("(_ \<Rightarrow>\<^sub>L /_)" [22, 21] 21) =
"{f::'a::real_normed_vector\<Rightarrow>'b::real_normed_vector. bounded_linear f}"
morphisms blinfun_apply Blinfun
by (blast intro: bounded_linear_intros)
@@ -142,7 +142,7 @@
instantiation blinfun :: (real_normed_vector, real_normed_vector) real_normed_vector
begin
-lift_definition%important norm_blinfun :: "'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> real" is onorm .
+lift_definition\<^marker>\<open>tag important\<close> norm_blinfun :: "'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> real" is onorm .
lift_definition minus_blinfun :: "'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b"
is "\<lambda>f g x. f x - g x"
@@ -160,14 +160,14 @@
lift_definition uminus_blinfun :: "'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b" is "\<lambda>f x. - f x"
by (rule bounded_linear_minus)
-lift_definition%important zero_blinfun :: "'a \<Rightarrow>\<^sub>L 'b" is "\<lambda>x. 0"
+lift_definition\<^marker>\<open>tag important\<close> zero_blinfun :: "'a \<Rightarrow>\<^sub>L 'b" is "\<lambda>x. 0"
by (rule bounded_linear_zero)
-lift_definition%important plus_blinfun :: "'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b"
+lift_definition\<^marker>\<open>tag important\<close> plus_blinfun :: "'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b"
is "\<lambda>f g x. f x + g x"
by (metis bounded_linear_add)
-lift_definition%important scaleR_blinfun::"real \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b" is "\<lambda>r f x. r *\<^sub>R f x"
+lift_definition\<^marker>\<open>tag important\<close> scaleR_blinfun::"real \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b" is "\<lambda>r f x. r *\<^sub>R f x"
by (metis bounded_linear_compose bounded_linear_scaleR_right)
definition sgn_blinfun :: "'a \<Rightarrow>\<^sub>L 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>L 'b"
@@ -367,7 +367,7 @@
by (rule convergentI)
qed
-subsection%unimportant \<open>On Euclidean Space\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>On Euclidean Space\<close>
lemma Zfun_sum:
assumes "finite s"
@@ -588,7 +588,7 @@
qed
-subsection%unimportant \<open>concrete bounded linear functions\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>concrete bounded linear functions\<close>
lemma transfer_bounded_bilinear_bounded_linearI:
assumes "g = (\<lambda>i x. (blinfun_apply (f i) x))"
@@ -757,7 +757,7 @@
defined analogously.
\<close>
-definition%important strong_operator_topology::"('a::real_normed_vector \<Rightarrow>\<^sub>L'b::real_normed_vector) topology"
+definition\<^marker>\<open>tag important\<close> strong_operator_topology::"('a::real_normed_vector \<Rightarrow>\<^sub>L'b::real_normed_vector) topology"
where "strong_operator_topology = pullback_topology UNIV blinfun_apply euclidean"
lemma strong_operator_topology_topspace:
--- a/src/HOL/Analysis/Brouwer_Fixpoint.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Brouwer_Fixpoint.thy Fri Apr 12 22:09:25 2019 +0200
@@ -231,16 +231,16 @@
definitions, but we need to split them into two implications because of the lack of type
quantifiers. Then ENR turns out to be equivalent to ANR plus local compactness."\<close>
-definition%important AR :: "'a::topological_space set \<Rightarrow> bool" where
+definition\<^marker>\<open>tag important\<close> AR :: "'a::topological_space set \<Rightarrow> bool" where
"AR S \<equiv> \<forall>U. \<forall>S'::('a * real) set.
S homeomorphic S' \<and> closedin (top_of_set U) S' \<longrightarrow> S' retract_of U"
-definition%important ANR :: "'a::topological_space set \<Rightarrow> bool" where
+definition\<^marker>\<open>tag important\<close> ANR :: "'a::topological_space set \<Rightarrow> bool" where
"ANR S \<equiv> \<forall>U. \<forall>S'::('a * real) set.
S homeomorphic S' \<and> closedin (top_of_set U) S'
\<longrightarrow> (\<exists>T. openin (top_of_set U) T \<and> S' retract_of T)"
-definition%important ENR :: "'a::topological_space set \<Rightarrow> bool" where
+definition\<^marker>\<open>tag important\<close> ENR :: "'a::topological_space set \<Rightarrow> bool" where
"ENR S \<equiv> \<exists>U. open U \<and> S retract_of U"
text \<open>First, show that we do indeed get the "usual" properties of ARs and ANRs.\<close>
@@ -4072,7 +4072,7 @@
by (metis finite_insert infinite_imp_nonempty less_linear sphere_eq_empty sphere_trivial)
qed
-corollary%unimportant ANR_sphere:
+corollary\<^marker>\<open>tag unimportant\<close> ANR_sphere:
fixes a :: "'a::euclidean_space"
shows "ANR(sphere a r)"
by (simp add: ENR_imp_ANR ENR_sphere)
@@ -4254,7 +4254,7 @@
qed
-corollary%unimportant nullhomotopic_into_ANR_extension:
+corollary\<^marker>\<open>tag unimportant\<close> nullhomotopic_into_ANR_extension:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "closed S"
and contf: "continuous_on S f"
@@ -4288,7 +4288,7 @@
by (force elim: homotopic_with_eq [of _ _ _ g "\<lambda>x. c"] simp: \<open>\<And>x. x \<in> S \<Longrightarrow> g x = f x\<close>)
qed
-corollary%unimportant nullhomotopic_into_rel_frontier_extension:
+corollary\<^marker>\<open>tag unimportant\<close> nullhomotopic_into_rel_frontier_extension:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "closed S"
and contf: "continuous_on S f"
@@ -4299,7 +4299,7 @@
(\<exists>g. continuous_on UNIV g \<and> range g \<subseteq> rel_frontier T \<and> (\<forall>x \<in> S. g x = f x))"
by (simp add: nullhomotopic_into_ANR_extension assms ANR_rel_frontier_convex)
-corollary%unimportant nullhomotopic_into_sphere_extension:
+corollary\<^marker>\<open>tag unimportant\<close> nullhomotopic_into_sphere_extension:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b :: euclidean_space"
assumes "closed S" and contf: "continuous_on S f"
and "S \<noteq> {}" and fim: "f ` S \<subseteq> sphere a r"
@@ -4321,7 +4321,7 @@
done
qed
-proposition%unimportant Borsuk_map_essential_bounded_component:
+proposition\<^marker>\<open>tag unimportant\<close> Borsuk_map_essential_bounded_component:
fixes a :: "'a :: euclidean_space"
assumes "compact S" and "a \<notin> S"
shows "bounded (connected_component_set (- S) a) \<longleftrightarrow>
@@ -4578,7 +4578,7 @@
qed
qed
-proposition%unimportant homotopic_neighbourhood_extension:
+proposition\<^marker>\<open>tag unimportant\<close> homotopic_neighbourhood_extension:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes contf: "continuous_on S f" and fim: "f ` S \<subseteq> U"
and contg: "continuous_on S g" and gim: "g ` S \<subseteq> U"
@@ -4643,7 +4643,7 @@
qed
text\<open> Homotopy on a union of closed-open sets.\<close>
-proposition%unimportant homotopic_on_clopen_Union:
+proposition\<^marker>\<open>tag unimportant\<close> homotopic_on_clopen_Union:
fixes \<F> :: "'a::euclidean_space set set"
assumes "\<And>S. S \<in> \<F> \<Longrightarrow> closedin (top_of_set (\<Union>\<F>)) S"
and "\<And>S. S \<in> \<F> \<Longrightarrow> openin (top_of_set (\<Union>\<F>)) S"
--- a/src/HOL/Analysis/Caratheodory.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Caratheodory.thy Fri Apr 12 22:09:25 2019 +0200
@@ -51,12 +51,12 @@
subsection \<open>Characterizations of Measures\<close>
-definition%important outer_measure_space where
+definition\<^marker>\<open>tag important\<close> outer_measure_space where
"outer_measure_space M f \<longleftrightarrow> positive M f \<and> increasing M f \<and> countably_subadditive M f"
-subsubsection%important \<open>Lambda Systems\<close>
+subsubsection\<^marker>\<open>tag important\<close> \<open>Lambda Systems\<close>
-definition%important lambda_system :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> 'a set set"
+definition\<^marker>\<open>tag important\<close> lambda_system :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> 'a set set"
where
"lambda_system \<Omega> M f = {l \<in> M. \<forall>x \<in> M. f (l \<inter> x) + f ((\<Omega> - l) \<inter> x) = f x}"
@@ -297,7 +297,7 @@
using pos by (simp add: measure_space_def)
qed
-definition%important outer_measure :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> 'a set \<Rightarrow> ennreal" where
+definition\<^marker>\<open>tag important\<close> outer_measure :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> 'a set \<Rightarrow> ennreal" where
"outer_measure M f X =
(INF A\<in>{A. range A \<subseteq> M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i)}. \<Sum>i. f (A i))"
@@ -499,7 +499,7 @@
subsection \<open>Volumes\<close>
-definition%important volume :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> bool" where
+definition\<^marker>\<open>tag important\<close> volume :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> bool" where
"volume M f \<longleftrightarrow>
(f {} = 0) \<and> (\<forall>a\<in>M. 0 \<le> f a) \<and>
(\<forall>C\<subseteq>M. disjoint C \<longrightarrow> finite C \<longrightarrow> \<Union>C \<in> M \<longrightarrow> f (\<Union>C) = (\<Sum>c\<in>C. f c))"
@@ -635,7 +635,7 @@
qed
qed
-subsubsection%important \<open>Caratheodory on semirings\<close>
+subsubsection\<^marker>\<open>tag important\<close> \<open>Caratheodory on semirings\<close>
theorem (in semiring_of_sets) caratheodory:
assumes pos: "positive M \<mu>" and ca: "countably_additive M \<mu>"
--- a/src/HOL/Analysis/Cartesian_Euclidean_Space.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Cartesian_Euclidean_Space.thy Fri Apr 12 22:09:25 2019 +0200
@@ -565,7 +565,7 @@
subsection "Derivative"
-definition%important "jacobian f net = matrix(frechet_derivative f net)"
+definition\<^marker>\<open>tag important\<close> "jacobian f net = matrix(frechet_derivative f net)"
proposition jacobian_works:
"(f::(real^'a) \<Rightarrow> (real^'b)) differentiable net \<longleftrightarrow>
@@ -591,7 +591,7 @@
vector_cart[of "\<lambda>j. frechet_derivative f (at x) j $ k"]
by (simp add: Basis_vec_def axis_eq_axis inner_axis jacobian_def matrix_def)
-subsection%unimportant\<open>Routine results connecting the types \<^typ>\<open>real^1\<close> and \<^typ>\<open>real\<close>\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Routine results connecting the types \<^typ>\<open>real^1\<close> and \<^typ>\<open>real\<close>\<close>
lemma vec_cbox_1_eq [simp]:
shows "vec ` cbox u v = cbox (vec u) (vec v ::real^1)"
--- a/src/HOL/Analysis/Cartesian_Space.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Cartesian_Space.thy Fri Apr 12 22:09:25 2019 +0200
@@ -13,7 +13,7 @@
Finite_Cartesian_Product Linear_Algebra
begin
-subsection%unimportant \<open>Type @{typ \<open>'a ^ 'n\<close>} and fields as vector spaces\<close> (*much of the following
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Type @{typ \<open>'a ^ 'n\<close>} and fields as vector spaces\<close> (*much of the following
is really basic linear algebra, check for overlap? rename subsection? *)
definition "cart_basis = {axis i 1 | i. i\<in>UNIV}"
@@ -552,7 +552,7 @@
using nullspace_inter_rowspace [of A "x-y"]
by (metis diff_eq_diff_eq diff_self matrix_vector_mult_diff_distrib span_diff)
-definition%important rank :: "'a::field^'n^'m=>nat"
+definition\<^marker>\<open>tag important\<close> rank :: "'a::field^'n^'m=>nat"
where row_rank_def_gen: "rank A \<equiv> vec.dim(rows A)"
lemma row_rank_def: "rank A = dim (rows A)" for A::"real^'n^'m"
@@ -671,7 +671,7 @@
by (metis matrix_transpose_mul rank_mul_le_right rank_transpose)
-subsection%unimportant \<open>Lemmas for working on \<open>real^1/2/3/4\<close>\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Lemmas for working on \<open>real^1/2/3/4\<close>\<close>
lemma exhaust_2:
fixes x :: 2
@@ -736,7 +736,7 @@
lemma sum_4: "sum f (UNIV::4 set) = f 1 + f 2 + f 3 + f 4"
unfolding UNIV_4 by (simp add: ac_simps)
-subsection%unimportant\<open>The collapse of the general concepts to dimension one\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>The collapse of the general concepts to dimension one\<close>
lemma vector_one: "(x::'a ^1) = (\<chi> i. (x$1))"
by (simp add: vec_eq_iff)
@@ -762,7 +762,7 @@
by (auto simp add: norm_real dist_norm)
-subsection%unimportant\<open>Routine results connecting the types \<^typ>\<open>real^1\<close> and \<^typ>\<open>real\<close>\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Routine results connecting the types \<^typ>\<open>real^1\<close> and \<^typ>\<open>real\<close>\<close>
lemma vector_one_nth [simp]:
fixes x :: "'a^1" shows "vec (x $ 1) = x"
@@ -795,7 +795,7 @@
done
-subsection%unimportant\<open>Explicit vector construction from lists\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Explicit vector construction from lists\<close>
definition "vector l = (\<chi> i. foldr (\<lambda>x f n. fun_upd (f (n+1)) n x) l (\<lambda>n x. 0) 1 i)"
@@ -835,7 +835,7 @@
apply (simp add: forall_3)
done
-subsection%unimportant \<open>lambda skolemization on cartesian products\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>lambda skolemization on cartesian products\<close>
lemma lambda_skolem: "(\<forall>i. \<exists>x. P i x) \<longleftrightarrow>
(\<exists>x::'a ^ 'n. \<forall>i. P i (x $ i))" (is "?lhs \<longleftrightarrow> ?rhs")
@@ -950,7 +950,7 @@
lemma const_vector_cart:"((\<chi> i. d)::real^'n) = (\<Sum>i\<in>Basis. d *\<^sub>R i)"
by (rule vector_cart)
-subsection%unimportant \<open>Explicit formulas for low dimensions\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Explicit formulas for low dimensions\<close>
lemma prod_neutral_const: "prod f {(1::nat)..1} = f 1"
by simp
@@ -964,7 +964,7 @@
subsection \<open>Orthogonality of a matrix\<close>
-definition%important "orthogonal_matrix (Q::'a::semiring_1^'n^'n) \<longleftrightarrow>
+definition\<^marker>\<open>tag important\<close> "orthogonal_matrix (Q::'a::semiring_1^'n^'n) \<longleftrightarrow>
transpose Q ** Q = mat 1 \<and> Q ** transpose Q = mat 1"
lemma orthogonal_matrix: "orthogonal_matrix (Q:: real ^'n^'n) \<longleftrightarrow> transpose Q ** Q = mat 1"
--- a/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Fri Apr 12 22:09:25 2019 +0200
@@ -6,7 +6,7 @@
imports Complex_Transcendental Weierstrass_Theorems Ordered_Euclidean_Space
begin
-subsection%unimportant \<open>Homeomorphisms of arc images\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Homeomorphisms of arc images\<close>
lemma homeomorphism_arc:
fixes g :: "real \<Rightarrow> 'a::t2_space"
@@ -85,7 +85,7 @@
using arc_simple_path inside_arc_empty by blast
-subsection%unimportant \<open>Piecewise differentiable functions\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Piecewise differentiable functions\<close>
definition piecewise_differentiable_on
(infixr "piecewise'_differentiable'_on" 50)
@@ -321,7 +321,7 @@
difficult proofs concerning winding numbers. We need a self-contained and straightforward theorem
asserting that all derivatives can be integrated before we can adopt Harrison's choice.\<close>
-definition%important C1_differentiable_on :: "(real \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> real set \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> C1_differentiable_on :: "(real \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> real set \<Rightarrow> bool"
(infix "C1'_differentiable'_on" 50)
where
"f C1_differentiable_on S \<longleftrightarrow>
@@ -400,7 +400,7 @@
by (rule continuous_intros | simp add: continuous_at_imp_continuous_on differentiable_imp_continuous_within)+
-definition%important piecewise_C1_differentiable_on
+definition\<^marker>\<open>tag important\<close> piecewise_C1_differentiable_on
(infixr "piecewise'_C1'_differentiable'_on" 50)
where "f piecewise_C1_differentiable_on i \<equiv>
continuous_on i f \<and>
@@ -704,7 +704,7 @@
subsection \<open>Valid paths, and their start and finish\<close>
-definition%important valid_path :: "(real \<Rightarrow> 'a :: real_normed_vector) \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> valid_path :: "(real \<Rightarrow> 'a :: real_normed_vector) \<Rightarrow> bool"
where "valid_path f \<equiv> f piecewise_C1_differentiable_on {0..1::real}"
definition closed_path :: "(real \<Rightarrow> 'a :: real_normed_vector) \<Rightarrow> bool"
@@ -803,17 +803,17 @@
text\<open>piecewise differentiable function on [0,1]\<close>
-definition%important has_contour_integral :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> (real \<Rightarrow> complex) \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> has_contour_integral :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> (real \<Rightarrow> complex) \<Rightarrow> bool"
(infixr "has'_contour'_integral" 50)
where "(f has_contour_integral i) g \<equiv>
((\<lambda>x. f(g x) * vector_derivative g (at x within {0..1}))
has_integral i) {0..1}"
-definition%important contour_integrable_on
+definition\<^marker>\<open>tag important\<close> contour_integrable_on
(infixr "contour'_integrable'_on" 50)
where "f contour_integrable_on g \<equiv> \<exists>i. (f has_contour_integral i) g"
-definition%important contour_integral
+definition\<^marker>\<open>tag important\<close> contour_integral
where "contour_integral g f \<equiv> SOME i. (f has_contour_integral i) g \<or> \<not> f contour_integrable_on g \<and> i=0"
lemma not_integrable_contour_integral: "\<not> f contour_integrable_on g \<Longrightarrow> contour_integral g f = 0"
@@ -870,7 +870,7 @@
(\<lambda>t. f(g t) * vector_derivative g (at t)) integrable_on {0..1}"
by (simp add: has_contour_integral integrable_on_def contour_integrable_on_def)
-subsection%unimportant \<open>Reversing a path\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Reversing a path\<close>
lemma valid_path_imp_reverse:
assumes "valid_path g"
@@ -948,7 +948,7 @@
qed
-subsection%unimportant \<open>Joining two paths together\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Joining two paths together\<close>
lemma valid_path_join:
assumes "valid_path g1" "valid_path g2" "pathfinish g1 = pathstart g2"
@@ -1136,7 +1136,7 @@
by (simp add: has_contour_integral_integral has_contour_integral_join contour_integral_unique)
-subsection%unimportant \<open>Shifting the starting point of a (closed) path\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Shifting the starting point of a (closed) path\<close>
lemma shiftpath_alt_def: "shiftpath a f = (\<lambda>x. if x \<le> 1-a then f (a + x) else f (a + x - 1))"
by (auto simp: shiftpath_def)
@@ -1279,7 +1279,7 @@
by (simp add: contour_integral_def contour_integrable_on_def has_contour_integral_shiftpath_eq)
-subsection%unimportant \<open>More about straight-line paths\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>More about straight-line paths\<close>
lemma has_vector_derivative_linepath_within:
"(linepath a b has_vector_derivative (b - a)) (at x within s)"
@@ -1669,7 +1669,7 @@
lemma contour_integrable_on_id [iff]: "(\<lambda>x. x) contour_integrable_on (linepath a b)"
by (simp add: continuous_on_id contour_integrable_continuous_linepath)
-subsection%unimportant \<open>Arithmetical combining theorems\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Arithmetical combining theorems\<close>
lemma has_contour_integral_neg:
"(f has_contour_integral i) g \<Longrightarrow> ((\<lambda>x. -(f x)) has_contour_integral (-i)) g"
@@ -1753,7 +1753,7 @@
\<Longrightarrow> ((\<lambda>x. sum (\<lambda>a. f a x) s) has_contour_integral sum i s) p"
by (induction s rule: finite_induct) (auto simp: has_contour_integral_0 has_contour_integral_add)
-subsection%unimportant \<open>Operations on path integrals\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Operations on path integrals\<close>
lemma contour_integral_const_linepath [simp]: "contour_integral (linepath a b) (\<lambda>x. c) = c*(b - a)"
by (rule contour_integral_unique [OF has_contour_integral_const_linepath])
@@ -1820,7 +1820,7 @@
by (metis has_contour_integral_eq)
-subsection%unimportant \<open>Arithmetic theorems for path integrability\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Arithmetic theorems for path integrability\<close>
lemma contour_integrable_neg:
"f contour_integrable_on g \<Longrightarrow> (\<lambda>x. -(f x)) contour_integrable_on g"
@@ -1858,7 +1858,7 @@
by (metis has_contour_integral_sum)
-subsection%unimportant \<open>Reversing a path integral\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Reversing a path integral\<close>
lemma has_contour_integral_reverse_linepath:
"(f has_contour_integral i) (linepath a b)
@@ -2088,7 +2088,7 @@
qed
-subsection%unimportant \<open>The key quadrisection step\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>The key quadrisection step\<close>
lemma norm_sum_half:
assumes "norm(a + b) \<ge> e"
@@ -2181,7 +2181,7 @@
qed
qed
-subsection%unimportant \<open>Cauchy's theorem for triangles\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Cauchy's theorem for triangles\<close>
lemma triangle_points_closer:
fixes a::complex
@@ -2304,7 +2304,7 @@
done
qed
-proposition%unimportant Cauchy_theorem_triangle:
+proposition\<^marker>\<open>tag unimportant\<close> Cauchy_theorem_triangle:
assumes "f holomorphic_on (convex hull {a,b,c})"
shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)"
proof -
@@ -2410,7 +2410,7 @@
using has_contour_integral_integral by fastforce
qed
-subsection%unimportant \<open>Version needing function holomorphic in interior only\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Version needing function holomorphic in interior only\<close>
lemma Cauchy_theorem_flat_lemma:
assumes f: "continuous_on (convex hull {a,b,c}) f"
@@ -2655,9 +2655,9 @@
using has_contour_integral_integral by fastforce
qed
-subsection%unimportant \<open>Version allowing finite number of exceptional points\<close>
-
-proposition%unimportant Cauchy_theorem_triangle_cofinite:
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Version allowing finite number of exceptional points\<close>
+
+proposition\<^marker>\<open>tag unimportant\<close> Cauchy_theorem_triangle_cofinite:
assumes "continuous_on (convex hull {a,b,c}) f"
and "finite S"
and "(\<And>x. x \<in> interior(convex hull {a,b,c}) - S \<Longrightarrow> f field_differentiable (at x))"
@@ -2740,7 +2740,7 @@
qed
qed
-subsection%unimportant \<open>Cauchy's theorem for an open starlike set\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Cauchy's theorem for an open starlike set\<close>
lemma starlike_convex_subset:
assumes S: "a \<in> S" "closed_segment b c \<subseteq> S" and subs: "\<And>x. x \<in> S \<Longrightarrow> closed_segment a x \<subseteq> S"
@@ -2957,7 +2957,7 @@
by (auto intro!: holomorphic_on_imp_differentiable_at simp: interior_open)
qed (use assms in \<open>auto intro: holomorphic_on_imp_continuous_on\<close>)
-corollary%unimportant Cauchy_theorem_convex:
+corollary\<^marker>\<open>tag unimportant\<close> Cauchy_theorem_convex:
"\<lbrakk>continuous_on S f; convex S; finite K;
\<And>x. x \<in> interior S - K \<Longrightarrow> f field_differentiable at x;
valid_path g; path_image g \<subseteq> S; pathfinish g = pathstart g\<rbrakk>
@@ -2973,19 +2973,19 @@
using at_within_interior holomorphic_on_def interior_subset by fastforce
text\<open>In particular for a disc\<close>
-corollary%unimportant Cauchy_theorem_disc:
+corollary\<^marker>\<open>tag unimportant\<close> Cauchy_theorem_disc:
"\<lbrakk>finite K; continuous_on (cball a e) f;
\<And>x. x \<in> ball a e - K \<Longrightarrow> f field_differentiable at x;
valid_path g; path_image g \<subseteq> cball a e;
pathfinish g = pathstart g\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) g"
by (auto intro: Cauchy_theorem_convex)
-corollary%unimportant Cauchy_theorem_disc_simple:
+corollary\<^marker>\<open>tag unimportant\<close> Cauchy_theorem_disc_simple:
"\<lbrakk>f holomorphic_on (ball a e); valid_path g; path_image g \<subseteq> ball a e;
pathfinish g = pathstart g\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) g"
by (simp add: Cauchy_theorem_convex_simple)
-subsection%unimportant \<open>Generalize integrability to local primitives\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Generalize integrability to local primitives\<close>
lemma contour_integral_local_primitive_lemma:
fixes f :: "complex\<Rightarrow>complex"
@@ -3422,7 +3422,7 @@
subsection \<open>Winding Numbers\<close>
-definition%important winding_number_prop :: "[real \<Rightarrow> complex, complex, real, real \<Rightarrow> complex, complex] \<Rightarrow> bool" where
+definition\<^marker>\<open>tag important\<close> winding_number_prop :: "[real \<Rightarrow> complex, complex, real, real \<Rightarrow> complex, complex] \<Rightarrow> bool" where
"winding_number_prop \<gamma> z e p n \<equiv>
valid_path p \<and> z \<notin> path_image p \<and>
pathstart p = pathstart \<gamma> \<and>
@@ -3430,7 +3430,7 @@
(\<forall>t \<in> {0..1}. norm(\<gamma> t - p t) < e) \<and>
contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * n"
-definition%important winding_number:: "[real \<Rightarrow> complex, complex] \<Rightarrow> complex" where
+definition\<^marker>\<open>tag important\<close> winding_number:: "[real \<Rightarrow> complex, complex] \<Rightarrow> complex" where
"winding_number \<gamma> z \<equiv> SOME n. \<forall>e > 0. \<exists>p. winding_number_prop \<gamma> z e p n"
@@ -3674,7 +3674,7 @@
done
qed
-subsubsection%unimportant \<open>Some lemmas about negating a path\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Some lemmas about negating a path\<close>
lemma valid_path_negatepath: "valid_path \<gamma> \<Longrightarrow> valid_path (uminus \<circ> \<gamma>)"
unfolding o_def using piecewise_C1_differentiable_neg valid_path_def by blast
@@ -3739,7 +3739,7 @@
by (simp add: valid_path_join path_image_join winding_number_join valid_path_imp_path)
-subsubsection%unimportant \<open>Useful sufficient conditions for the winding number to be positive\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Useful sufficient conditions for the winding number to be positive\<close>
lemma Re_winding_number:
"\<lbrakk>valid_path \<gamma>; z \<notin> path_image \<gamma>\<rbrakk>
@@ -4228,7 +4228,7 @@
"path \<gamma> \<Longrightarrow> continuous_on (- path_image \<gamma>) (\<lambda>w. winding_number \<gamma> w)"
by (simp add: continuous_at_imp_continuous_on continuous_at_winding_number)
-subsection%unimportant \<open>The winding number is constant on a connected region\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>The winding number is constant on a connected region\<close>
lemma winding_number_constant:
assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>" and cs: "connected S" and sg: "S \<inter> path_image \<gamma> = {}"
@@ -4329,11 +4329,11 @@
finally show ?thesis .
qed
-corollary%unimportant winding_number_zero_const: "a \<noteq> z \<Longrightarrow> winding_number (\<lambda>t. a) z = 0"
+corollary\<^marker>\<open>tag unimportant\<close> winding_number_zero_const: "a \<noteq> z \<Longrightarrow> winding_number (\<lambda>t. a) z = 0"
by (rule winding_number_zero_in_outside)
(auto simp: pathfinish_def pathstart_def path_polynomial_function)
-corollary%unimportant winding_number_zero_outside:
+corollary\<^marker>\<open>tag unimportant\<close> winding_number_zero_outside:
"\<lbrakk>path \<gamma>; convex s; pathfinish \<gamma> = pathstart \<gamma>; z \<notin> s; path_image \<gamma> \<subseteq> s\<rbrakk> \<Longrightarrow> winding_number \<gamma> z = 0"
by (meson convex_in_outside outside_mono subsetCE winding_number_zero_in_outside)
@@ -4791,7 +4791,7 @@
apply (blast dest: holomorphic_on_imp_continuous_on homotopic_loops_imp_subset)
by (simp add: Cauchy_theorem_homotopic_loops)
-subsection%unimportant \<open>More winding number properties\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>More winding number properties\<close>
text\<open>including the fact that it's +-1 inside a simple closed curve.\<close>
@@ -4908,7 +4908,7 @@
subsection\<open>Partial circle path\<close>
-definition%important part_circlepath :: "[complex, real, real, real, real] \<Rightarrow> complex"
+definition\<^marker>\<open>tag important\<close> part_circlepath :: "[complex, real, real, real, real] \<Rightarrow> complex"
where "part_circlepath z r s t \<equiv> \<lambda>x. z + of_real r * exp (\<i> * of_real (linepath s t x))"
lemma pathstart_part_circlepath [simp]:
@@ -5294,7 +5294,7 @@
subsection\<open>Special case of one complete circle\<close>
-definition%important circlepath :: "[complex, real, real] \<Rightarrow> complex"
+definition\<^marker>\<open>tag important\<close> circlepath :: "[complex, real, real] \<Rightarrow> complex"
where "circlepath z r \<equiv> part_circlepath z r 0 (2*pi)"
lemma circlepath: "circlepath z r = (\<lambda>x. z + r * exp(2 * of_real pi * \<i> * of_real x))"
@@ -5487,7 +5487,7 @@
by (simp add: winding_number_circlepath assms)
qed
-corollary%unimportant Cauchy_integral_circlepath_simple:
+corollary\<^marker>\<open>tag unimportant\<close> Cauchy_integral_circlepath_simple:
assumes "f holomorphic_on cball z r" "norm(w - z) < r"
shows "((\<lambda>u. f u/(u - w)) has_contour_integral (2 * of_real pi * \<i> * f w))
(circlepath z r)"
@@ -5584,7 +5584,7 @@
by (rule tendstoI)
qed
-corollary%unimportant contour_integral_uniform_limit_circlepath:
+corollary\<^marker>\<open>tag unimportant\<close> 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 "\<not> trivial_limit F" "0 < r"
@@ -5593,7 +5593,7 @@
using assms by (auto simp: vector_derivative_circlepath norm_mult intro!: contour_integral_uniform_limit)
-subsection%unimportant \<open>General stepping result for derivative formulas\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>General stepping result for derivative formulas\<close>
lemma Cauchy_next_derivative:
assumes "continuous_on (path_image \<gamma>) f'"
@@ -6696,7 +6696,7 @@
qed
-subsection%unimportant \<open>Some more simple/convenient versions for applications\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Some more simple/convenient versions for applications\<close>
lemma holomorphic_uniform_sequence:
assumes S: "open S"
@@ -6896,7 +6896,7 @@
using less_le_trans norm_not_less_zero by blast
qed
-proposition%unimportant power_series_and_derivative:
+proposition\<^marker>\<open>tag unimportant\<close> power_series_and_derivative:
fixes a :: "nat \<Rightarrow> complex" and r::real
assumes "summable (\<lambda>n. a n * r^n)"
obtains g g' where "\<forall>z \<in> ball w r.
@@ -6909,7 +6909,7 @@
apply (auto simp: norm_minus_commute Ball_def dist_norm)
done
-proposition%unimportant power_series_holomorphic:
+proposition\<^marker>\<open>tag unimportant\<close> power_series_holomorphic:
assumes "\<And>w. w \<in> ball z r \<Longrightarrow> ((\<lambda>n. a n*(w - z)^n) sums f w)"
shows "f holomorphic_on ball z r"
proof -
@@ -6965,7 +6965,7 @@
(\<forall>w \<in> ball z r. (\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)"
by (simp add: analytic_on_open holomorphic_iff_power_series)
-subsection%unimportant \<open>Equality between holomorphic functions, on open ball then connected set\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Equality between holomorphic functions, on open ball then connected set\<close>
lemma holomorphic_fun_eq_on_ball:
"\<lbrakk>f holomorphic_on ball z r; g holomorphic_on ball z r;
@@ -7048,7 +7048,7 @@
by (subst higher_deriv_diff) (use assms in \<open>auto intro: holomorphic_intros\<close>)
qed (use assms in auto)
-subsection%unimportant \<open>Some basic lemmas about poles/singularities\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Some basic lemmas about poles/singularities\<close>
lemma pole_lemma:
assumes holf: "f holomorphic_on S" and a: "a \<in> interior S"
@@ -7775,7 +7775,7 @@
homotopic_paths_imp_homotopic_loops)
using valid_path_imp_path by blast
-proposition%unimportant holomorphic_logarithm_exists:
+proposition\<^marker>\<open>tag unimportant\<close> holomorphic_logarithm_exists:
assumes A: "convex A" "open A"
and f: "f holomorphic_on A" "\<And>x. x \<in> A \<Longrightarrow> f x \<noteq> 0"
and z0: "z0 \<in> A"
--- a/src/HOL/Analysis/Change_Of_Vars.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Change_Of_Vars.thy Fri Apr 12 22:09:25 2019 +0200
@@ -424,10 +424,10 @@
subsection\<open>\<open>F_sigma\<close> and \<open>G_delta\<close> sets.\<close>(*FIX ME mv *)
(*https://en.wikipedia.org/wiki/F\<sigma>_set*)
-inductive%important fsigma :: "'a::topological_space set \<Rightarrow> bool" where
+inductive\<^marker>\<open>tag important\<close> fsigma :: "'a::topological_space set \<Rightarrow> bool" where
"(\<And>n::nat. closed (F n)) \<Longrightarrow> fsigma (\<Union>(F ` UNIV))"
-inductive%important gdelta :: "'a::topological_space set \<Rightarrow> bool" where
+inductive\<^marker>\<open>tag important\<close> gdelta :: "'a::topological_space set \<Rightarrow> bool" where
"(\<And>n::nat. open (F n)) \<Longrightarrow> gdelta (\<Inter>(F ` UNIV))"
lemma fsigma_Union_compact:
@@ -2100,7 +2100,7 @@
assumes S: "S \<in> sets lebesgue" and f: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
shows "(\<lambda>x. det(matrix(f' x))) \<in> borel_measurable (lebesgue_on S)"
unfolding det_def
- by%unimportant (intro measurable) (auto intro: f borel_measurable_partial_derivatives [OF S])
+ by (intro measurable) (auto intro: f borel_measurable_partial_derivatives [OF S])
text\<open>The localisation wrt S uses the same argument for many similar results.\<close>
(*See HOL Light's MEASURABLE_ON_LEBESGUE_MEASURABLE_PREIMAGE_BOREL, etc.*)
@@ -2949,7 +2949,7 @@
and intS: "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S"
shows "f absolutely_integrable_on (g ` S)"
apply (rule absolutely_integrable_componentwise [OF absolutely_integrable_on_image_real [OF der_g]])
- using%unimportant absolutely_integrable_component [OF intS] by%unimportant auto
+ using absolutely_integrable_component [OF intS] by auto
proposition integral_on_image_ubound:
fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
@@ -2957,7 +2957,7 @@
and "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
and "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) integrable_on S"
shows "integral (g ` S) f \<le> integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x))"
- using%unimportant integral_on_image_ubound_nonneg [OF assms] by%unimportant simp
+ using integral_on_image_ubound_nonneg [OF assms] by simp
subsection\<open>Change-of-variables theorem\<close>
@@ -3576,7 +3576,7 @@
and "inj_on g S"
shows "f absolutely_integrable_on (g ` S)
\<longleftrightarrow> (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S"
- using%unimportant assms has_absolute_integral_change_of_variables by%unimportant blast
+ using assms has_absolute_integral_change_of_variables by blast
corollary integral_change_of_variables:
fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
@@ -3586,8 +3586,8 @@
and disj: "(f absolutely_integrable_on (g ` S) \<or>
(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S)"
shows "integral (g ` S) f = integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x))"
- using%unimportant has_absolute_integral_change_of_variables [OF S der_g inj] disj
- by%unimportant blast
+ using has_absolute_integral_change_of_variables [OF S der_g inj] disj
+ by blast
lemma has_absolute_integral_change_of_variables_1:
fixes f :: "real \<Rightarrow> real^'n::{finite,wellorder}" and g :: "real \<Rightarrow> real"
@@ -3635,7 +3635,7 @@
and inj: "inj_on g S"
shows "(f absolutely_integrable_on g ` S \<longleftrightarrow>
(\<lambda>x. \<bar>g' x\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S)"
- using%unimportant has_absolute_integral_change_of_variables_1 [OF assms] by%unimportant auto
+ using has_absolute_integral_change_of_variables_1 [OF assms] by auto
subsection\<open>Change of variables for integrals: special case of linear function\<close>
--- a/src/HOL/Analysis/Complete_Measure.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Complete_Measure.thy Fri Apr 12 22:09:25 2019 +0200
@@ -6,21 +6,21 @@
imports Bochner_Integration
begin
-locale%important complete_measure =
+locale\<^marker>\<open>tag important\<close> complete_measure =
fixes M :: "'a measure"
assumes complete: "\<And>A B. B \<subseteq> A \<Longrightarrow> A \<in> null_sets M \<Longrightarrow> B \<in> sets M"
-definition%important
+definition\<^marker>\<open>tag important\<close>
"split_completion M A p = (if A \<in> sets M then p = (A, {}) else
\<exists>N'. A = fst p \<union> snd p \<and> fst p \<inter> snd p = {} \<and> fst p \<in> sets M \<and> snd p \<subseteq> N' \<and> N' \<in> null_sets M)"
-definition%important
+definition\<^marker>\<open>tag important\<close>
"main_part M A = fst (Eps (split_completion M A))"
-definition%important
+definition\<^marker>\<open>tag important\<close>
"null_part M A = snd (Eps (split_completion M A))"
-definition%important completion :: "'a measure \<Rightarrow> 'a measure" where
+definition\<^marker>\<open>tag important\<close> completion :: "'a measure \<Rightarrow> 'a measure" where
"completion M = measure_of (space M) { S \<union> N |S N N'. S \<in> sets M \<and> N' \<in> null_sets M \<and> N \<subseteq> N' }
(emeasure M \<circ> main_part M)"
@@ -67,9 +67,9 @@
by (intro completionI[of _ "\<Union>(S ` UNIV)" "\<Union>(N ` UNIV)" "\<Union>(N' ` UNIV)"]) auto
qed
-lemma%important sets_completion:
+lemma\<^marker>\<open>tag important\<close> sets_completion:
"sets (completion M) = { S \<union> N |S N N'. S \<in> sets M \<and> N' \<in> null_sets M \<and> N \<subseteq> N' }"
- using%unimportant sigma_algebra.sets_measure_of_eq[OF sigma_algebra_completion]
+ using sigma_algebra.sets_measure_of_eq[OF sigma_algebra_completion]
by (simp add: completion_def)
lemma sets_completionE:
@@ -86,17 +86,17 @@
"A \<in> sets M \<Longrightarrow> A \<in> sets (completion M)"
unfolding sets_completion by force
-lemma%important measurable_completion: "f \<in> M \<rightarrow>\<^sub>M N \<Longrightarrow> f \<in> completion M \<rightarrow>\<^sub>M N"
- by%unimportant (auto simp: measurable_def)
+lemma\<^marker>\<open>tag important\<close> measurable_completion: "f \<in> M \<rightarrow>\<^sub>M N \<Longrightarrow> f \<in> completion M \<rightarrow>\<^sub>M N"
+ by (auto simp: measurable_def)
lemma null_sets_completion:
assumes "N' \<in> null_sets M" "N \<subseteq> N'" shows "N \<in> sets (completion M)"
using assms by (intro sets_completionI[of N "{}" N N']) auto
-lemma%important split_completion:
+lemma\<^marker>\<open>tag important\<close> split_completion:
assumes "A \<in> sets (completion M)"
shows "split_completion M A (main_part M A, null_part M A)"
-proof%unimportant cases
+proof cases
assume "A \<in> sets M" then show ?thesis
by (simp add: split_completion_def[abs_def] main_part_def null_part_def)
next
@@ -181,10 +181,10 @@
finally show ?thesis .
qed
-lemma%important emeasure_completion[simp]:
+lemma\<^marker>\<open>tag important\<close> emeasure_completion[simp]:
assumes S: "S \<in> sets (completion M)"
shows "emeasure (completion M) S = emeasure M (main_part M S)"
-proof%unimportant (subst emeasure_measure_of[OF completion_def completion_into_space])
+proof (subst emeasure_measure_of[OF completion_def completion_into_space])
let ?\<mu> = "emeasure M \<circ> main_part M"
show "S \<in> sets (completion M)" "?\<mu> S = emeasure M (main_part M S) " using S by simp_all
show "positive (sets (completion M)) ?\<mu>"
@@ -283,11 +283,11 @@
qed
qed
-lemma%important completion_ex_borel_measurable:
+lemma\<^marker>\<open>tag important\<close> completion_ex_borel_measurable:
fixes g :: "'a \<Rightarrow> ennreal"
assumes g: "g \<in> borel_measurable (completion M)"
shows "\<exists>g'\<in>borel_measurable M. (AE x in M. g x = g' x)"
-proof%unimportant -
+proof -
from g[THEN borel_measurable_implies_simple_function_sequence'] guess f . note f = this
from this(1)[THEN completion_ex_simple_function]
have "\<forall>i. \<exists>f'. simple_function M f' \<and> (AE x in M. f i x = f' x)" ..
@@ -442,19 +442,19 @@
assumes f: "f \<in> M \<rightarrow>\<^sub>M borel" shows "integral\<^sup>L (completion M) f = integral\<^sup>L M f"
by (rule integral_subalgebra[symmetric]) (auto intro: f)
-locale%important semifinite_measure =
+locale\<^marker>\<open>tag important\<close> semifinite_measure =
fixes M :: "'a measure"
assumes semifinite:
"\<And>A. A \<in> sets M \<Longrightarrow> emeasure M A = \<infinity> \<Longrightarrow> \<exists>B\<in>sets M. B \<subseteq> A \<and> emeasure M B < \<infinity>"
-locale%important locally_determined_measure = semifinite_measure +
+locale\<^marker>\<open>tag important\<close> locally_determined_measure = semifinite_measure +
assumes locally_determined:
"\<And>A. A \<subseteq> space M \<Longrightarrow> (\<And>B. B \<in> sets M \<Longrightarrow> emeasure M B < \<infinity> \<Longrightarrow> A \<inter> B \<in> sets M) \<Longrightarrow> A \<in> sets M"
-locale%important cld_measure =
+locale\<^marker>\<open>tag important\<close> cld_measure =
complete_measure M + locally_determined_measure M for M :: "'a measure"
-definition%important outer_measure_of :: "'a measure \<Rightarrow> 'a set \<Rightarrow> ennreal"
+definition\<^marker>\<open>tag important\<close> outer_measure_of :: "'a measure \<Rightarrow> 'a set \<Rightarrow> ennreal"
where "outer_measure_of M A = (INF B \<in> {B\<in>sets M. A \<subseteq> B}. emeasure M B)"
lemma outer_measure_of_eq[simp]: "A \<in> sets M \<Longrightarrow> outer_measure_of M A = emeasure M A"
@@ -553,7 +553,7 @@
by (simp add: eq F)
qed (auto intro: SUP_least outer_measure_of_mono)
-definition%important measurable_envelope :: "'a measure \<Rightarrow> 'a set \<Rightarrow> 'a set \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> measurable_envelope :: "'a measure \<Rightarrow> 'a set \<Rightarrow> 'a set \<Rightarrow> bool"
where "measurable_envelope M A E \<longleftrightarrow>
(A \<subseteq> E \<and> E \<in> sets M \<and> (\<forall>F\<in>sets M. emeasure M (F \<inter> E) = outer_measure_of M (F \<inter> A)))"
@@ -618,7 +618,7 @@
by (auto simp: Int_absorb1)
qed
-lemma%important measurable_envelope_eq2:
+lemma\<^marker>\<open>tag important\<close> measurable_envelope_eq2:
assumes "A \<subseteq> E" "E \<in> sets M" "emeasure M E < \<infinity>"
shows "measurable_envelope M A E \<longleftrightarrow> (emeasure M E = outer_measure_of M A)"
proof safe
--- a/src/HOL/Analysis/Complex_Analysis_Basics.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Complex_Analysis_Basics.thy Fri Apr 12 22:09:25 2019 +0200
@@ -10,7 +10,7 @@
(* TODO FIXME: A lot of the things in here have nothing to do with complex analysis *)
-subsection%unimportant\<open>General lemmas\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>General lemmas\<close>
lemma nonneg_Reals_cmod_eq_Re: "z \<in> \<real>\<^sub>\<ge>\<^sub>0 \<Longrightarrow> norm z = Re z"
by (simp add: complex_nonneg_Reals_iff cmod_eq_Re)
@@ -236,11 +236,11 @@
subsection\<open>Holomorphic functions\<close>
-definition%important holomorphic_on :: "[complex \<Rightarrow> complex, complex set] \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> holomorphic_on :: "[complex \<Rightarrow> complex, complex set] \<Rightarrow> bool"
(infixl "(holomorphic'_on)" 50)
where "f holomorphic_on s \<equiv> \<forall>x\<in>s. f field_differentiable (at x within s)"
-named_theorems%important holomorphic_intros "structural introduction rules for holomorphic_on"
+named_theorems\<^marker>\<open>tag important\<close> holomorphic_intros "structural introduction rules for holomorphic_on"
lemma holomorphic_onI [intro?]: "(\<And>x. x \<in> s \<Longrightarrow> f field_differentiable (at x within s)) \<Longrightarrow> f holomorphic_on s"
by (simp add: holomorphic_on_def)
@@ -498,7 +498,7 @@
by (rule nonzero_deriv_nonconstant [of f "deriv f \<xi>" \<xi> S])
(use assms in \<open>auto simp: holomorphic_derivI\<close>)
-subsection%unimportant\<open>Caratheodory characterization\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Caratheodory characterization\<close>
lemma field_differentiable_caratheodory_at:
"f field_differentiable (at z) \<longleftrightarrow>
@@ -514,10 +514,10 @@
subsection\<open>Analyticity on a set\<close>
-definition%important analytic_on (infixl "(analytic'_on)" 50)
+definition\<^marker>\<open>tag important\<close> analytic_on (infixl "(analytic'_on)" 50)
where "f analytic_on S \<equiv> \<forall>x \<in> S. \<exists>e. 0 < e \<and> f holomorphic_on (ball x e)"
-named_theorems%important analytic_intros "introduction rules for proving analyticity"
+named_theorems\<^marker>\<open>tag important\<close> analytic_intros "introduction rules for proving analyticity"
lemma analytic_imp_holomorphic: "f analytic_on S \<Longrightarrow> f holomorphic_on S"
by (simp add: at_within_open [OF _ open_ball] analytic_on_def holomorphic_on_def)
@@ -734,7 +734,7 @@
finally show ?thesis .
qed
-subsection%unimportant\<open>Analyticity at a point\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Analyticity at a point\<close>
lemma analytic_at_ball:
"f analytic_on {z} \<longleftrightarrow> (\<exists>e. 0<e \<and> f holomorphic_on ball z e)"
@@ -769,7 +769,7 @@
by (force simp add: analytic_at)
qed
-subsection%unimportant\<open>Combining theorems for derivative with ``analytic at'' hypotheses\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Combining theorems for derivative with ``analytic at'' hypotheses\<close>
lemma
assumes "f analytic_on {z}" "g analytic_on {z}"
@@ -805,7 +805,7 @@
"f analytic_on {z} \<Longrightarrow> deriv (\<lambda>w. f w * c) z = deriv f z * c"
by (auto simp: complex_derivative_mult_at deriv_const analytic_on_const)
-subsection%unimportant\<open>Complex differentiation of sequences and series\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Complex differentiation of sequences and series\<close>
(* TODO: Could probably be simplified using Uniform_Limit *)
lemma has_complex_derivative_sequence:
@@ -910,7 +910,7 @@
by (auto simp: summable_def field_differentiable_def has_field_derivative_def)
qed
-subsection%unimportant\<open>Bound theorem\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Bound theorem\<close>
lemma field_differentiable_bound:
fixes S :: "'a::real_normed_field set"
@@ -924,7 +924,7 @@
apply (rule onorm_le, simp_all add: norm_mult mult_right_mono assms)
done
-subsection%unimportant\<open>Inverse function theorem for complex derivatives\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Inverse function theorem for complex derivatives\<close>
lemma has_field_derivative_inverse_basic:
shows "DERIV f (g y) :> f' \<Longrightarrow>
@@ -965,7 +965,7 @@
apply (rule has_derivative_inverse_strong_x [of S g y f])
by auto
-subsection%unimportant \<open>Taylor on Complex Numbers\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Taylor on Complex Numbers\<close>
lemma sum_Suc_reindex:
fixes f :: "nat \<Rightarrow> 'a::ab_group_add"
@@ -1151,7 +1151,7 @@
qed
-subsection%unimportant \<open>Polynomal function extremal theorem, from HOL Light\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Polynomal function extremal theorem, from HOL Light\<close>
lemma polyfun_extremal_lemma: (*COMPLEX_POLYFUN_EXTREMAL_LEMMA in HOL Light*)
fixes c :: "nat \<Rightarrow> 'a::real_normed_div_algebra"
--- a/src/HOL/Analysis/Complex_Transcendental.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Complex_Transcendental.thy Fri Apr 12 22:09:25 2019 +0200
@@ -12,7 +12,7 @@
subsection\<open>Möbius transformations\<close>
(* TODO: Figure out what to do with Möbius transformations *)
-definition%important "moebius a b c d = (\<lambda>z. (a*z+b) / (c*z+d :: 'a :: field))"
+definition\<^marker>\<open>tag important\<close> "moebius a b c d = (\<lambda>z. (a*z+b) / (c*z+d :: 'a :: field))"
theorem moebius_inverse:
assumes "a * d \<noteq> b * c" "c * z + d \<noteq> 0"
@@ -64,7 +64,7 @@
by (simp add: norm_power Im_power2)
qed
-subsection%unimportant\<open>The Exponential Function\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>The Exponential Function\<close>
lemma norm_exp_i_times [simp]: "norm (exp(\<i> * of_real y)) = 1"
by simp
@@ -116,7 +116,7 @@
using sums_unique2 by blast
qed
-corollary%unimportant exp_minus_Euler: "exp(-(\<i> * z)) = cos(z) - \<i> * sin(z)"
+corollary\<^marker>\<open>tag unimportant\<close> exp_minus_Euler: "exp(-(\<i> * z)) = cos(z) - \<i> * sin(z)"
using exp_Euler [of "-z"]
by simp
@@ -154,7 +154,7 @@
lemma Im_sin_nonneg2: "Re z = pi \<Longrightarrow> Im z \<le> 0 \<Longrightarrow> 0 \<le> Im (sin z)"
by (simp add: Re_sin Im_sin algebra_simps)
-subsection%unimportant\<open>Relationships between real and complex trigonometric and hyperbolic functions\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Relationships between real and complex trigonometric and hyperbolic functions\<close>
lemma real_sin_eq [simp]: "Re(sin(of_real x)) = sin x"
by (simp add: sin_of_real)
@@ -213,7 +213,7 @@
shows "(\<lambda>x. cos (f x)) holomorphic_on A"
using holomorphic_on_compose[OF assms holomorphic_on_cos] by (simp add: o_def)
-subsection%unimportant\<open>More on the Polar Representation of Complex Numbers\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>More on the Polar Representation of Complex Numbers\<close>
lemma exp_Complex: "exp(Complex r t) = of_real(exp r) * Complex (cos t) (sin t)"
by (simp add: Complex_eq exp_add exp_Euler exp_of_real sin_of_real cos_of_real)
@@ -693,7 +693,7 @@
qed
-subsection%unimportant\<open>Taylor series for complex exponential, sine and cosine\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Taylor series for complex exponential, sine and cosine\<close>
declare power_Suc [simp del]
@@ -859,10 +859,10 @@
subsection\<open>The argument of a complex number (HOL Light version)\<close>
-definition%important is_Arg :: "[complex,real] \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> is_Arg :: "[complex,real] \<Rightarrow> bool"
where "is_Arg z r \<equiv> z = of_real(norm z) * exp(\<i> * of_real r)"
-definition%important Arg2pi :: "complex \<Rightarrow> real"
+definition\<^marker>\<open>tag important\<close> Arg2pi :: "complex \<Rightarrow> real"
where "Arg2pi z \<equiv> if z = 0 then 0 else THE t. 0 \<le> t \<and> t < 2*pi \<and> is_Arg z t"
lemma is_Arg_2pi_iff: "is_Arg z (r + of_int k * (2 * pi)) \<longleftrightarrow> is_Arg z r"
@@ -940,7 +940,7 @@
done
qed
-corollary%unimportant
+corollary\<^marker>\<open>tag unimportant\<close>
shows Arg2pi_ge_0: "0 \<le> Arg2pi z"
and Arg2pi_lt_2pi: "Arg2pi z < 2*pi"
and Arg2pi_eq: "z = of_real(norm z) * exp(\<i> * of_real(Arg2pi z))"
@@ -1023,7 +1023,7 @@
by blast
qed auto
-corollary%unimportant Arg2pi_gt_0:
+corollary\<^marker>\<open>tag unimportant\<close> Arg2pi_gt_0:
assumes "z \<notin> \<real>\<^sub>\<ge>\<^sub>0"
shows "Arg2pi z > 0"
using Arg2pi_eq_0 Arg2pi_ge_0 assms dual_order.strict_iff_order
@@ -1128,7 +1128,7 @@
by (metis (no_types, hide_lams) abs_le_D1 distrib_left mult.commute mult.left_commute mult_left_le)
qed
-subsection%unimportant\<open>Analytic properties of tangent function\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Analytic properties of tangent function\<close>
lemma cnj_tan: "cnj(tan z) = tan(cnj z)"
by (simp add: cnj_cos cnj_sin tan_def)
@@ -1156,7 +1156,7 @@
instantiation complex :: ln
begin
-definition%important ln_complex :: "complex \<Rightarrow> complex"
+definition\<^marker>\<open>tag important\<close> ln_complex :: "complex \<Rightarrow> complex"
where "ln_complex \<equiv> \<lambda>z. THE w. exp w = z & -pi < Im(w) & Im(w) \<le> pi"
text\<open>NOTE: within this scope, the constant Ln is not yet available!\<close>
@@ -1189,7 +1189,7 @@
apply auto
done
-subsection%unimportant\<open>Relation to Real Logarithm\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Relation to Real Logarithm\<close>
lemma Ln_of_real:
assumes "0 < z"
@@ -1203,10 +1203,10 @@
using assms by simp
qed
-corollary%unimportant Ln_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> Re z > 0 \<Longrightarrow> ln z \<in> \<real>"
+corollary\<^marker>\<open>tag unimportant\<close> Ln_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> Re z > 0 \<Longrightarrow> ln z \<in> \<real>"
by (auto simp: Ln_of_real elim: Reals_cases)
-corollary%unimportant Im_Ln_of_real [simp]: "r > 0 \<Longrightarrow> Im (ln (of_real r)) = 0"
+corollary\<^marker>\<open>tag unimportant\<close> Im_Ln_of_real [simp]: "r > 0 \<Longrightarrow> Im (ln (of_real r)) = 0"
by (simp add: Ln_of_real)
lemma cmod_Ln_Reals [simp]: "z \<in> \<real> \<Longrightarrow> 0 < Re z \<Longrightarrow> cmod (ln z) = norm (ln (Re z))"
@@ -1244,13 +1244,13 @@
lemma Re_Ln [simp]: "z \<noteq> 0 \<Longrightarrow> Re(Ln z) = ln(norm z)"
by (metis exp_Ln ln_exp norm_exp_eq_Re)
-corollary%unimportant ln_cmod_le:
+corollary\<^marker>\<open>tag unimportant\<close> ln_cmod_le:
assumes z: "z \<noteq> 0"
shows "ln (cmod z) \<le> cmod (Ln z)"
using norm_exp [of "Ln z", simplified exp_Ln [OF z]]
by (metis Re_Ln complex_Re_le_cmod z)
-proposition%unimportant exists_complex_root:
+proposition\<^marker>\<open>tag unimportant\<close> exists_complex_root:
fixes z :: complex
assumes "n \<noteq> 0" obtains w where "z = w ^ n"
proof (cases "z=0")
@@ -1259,13 +1259,13 @@
by (rule_tac w = "exp(Ln z / n)" in that) (simp add: assms exp_of_nat_mult [symmetric])
qed (use assms in auto)
-corollary%unimportant exists_complex_root_nonzero:
+corollary\<^marker>\<open>tag unimportant\<close> exists_complex_root_nonzero:
fixes z::complex
assumes "z \<noteq> 0" "n \<noteq> 0"
obtains w where "w \<noteq> 0" "z = w ^ n"
by (metis exists_complex_root [of n z] assms power_0_left)
-subsection%unimportant\<open>Derivative of Ln away from the branch cut\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Derivative of Ln away from the branch cut\<close>
lemma
assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0"
@@ -1456,7 +1456,7 @@
qed
-subsection%unimportant\<open>Quadrant-type results for Ln\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Quadrant-type results for Ln\<close>
lemma cos_lt_zero_pi: "pi/2 < x \<Longrightarrow> x < 3*pi/2 \<Longrightarrow> cos x < 0"
using cos_minus_pi cos_gt_zero_pi [of "x-pi"]
@@ -1553,7 +1553,7 @@
mult_zero_right not_less_iff_gr_or_eq pi_ge_zero pi_neq_zero rcis_zero_arg rcis_zero_mod)
-subsection%unimportant\<open>More Properties of Ln\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>More Properties of Ln\<close>
lemma cnj_Ln: assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0" shows "cnj(Ln z) = Ln(cnj z)"
proof (cases "z=0")
@@ -1631,27 +1631,27 @@
using assms mpi_less_Im_Ln [of w] mpi_less_Im_Ln [of z]
by (auto simp: exp_add exp_diff sin_double cos_double exp_Euler intro!: Ln_unique)
-corollary%unimportant Ln_times_simple:
+corollary\<^marker>\<open>tag unimportant\<close> Ln_times_simple:
"\<lbrakk>w \<noteq> 0; z \<noteq> 0; -pi < Im(Ln w) + Im(Ln z); Im(Ln w) + Im(Ln z) \<le> pi\<rbrakk>
\<Longrightarrow> Ln(w * z) = Ln(w) + Ln(z)"
by (simp add: Ln_times)
-corollary%unimportant Ln_times_of_real:
+corollary\<^marker>\<open>tag unimportant\<close> Ln_times_of_real:
"\<lbrakk>r > 0; z \<noteq> 0\<rbrakk> \<Longrightarrow> Ln(of_real r * z) = ln r + Ln(z)"
using mpi_less_Im_Ln Im_Ln_le_pi
by (force simp: Ln_times)
-corollary%unimportant Ln_times_Reals:
+corollary\<^marker>\<open>tag unimportant\<close> Ln_times_Reals:
"\<lbrakk>r \<in> Reals; Re r > 0; z \<noteq> 0\<rbrakk> \<Longrightarrow> Ln(r * z) = ln (Re r) + Ln(z)"
using Ln_Reals_eq Ln_times_of_real by fastforce
-corollary%unimportant Ln_divide_of_real:
+corollary\<^marker>\<open>tag unimportant\<close> Ln_divide_of_real:
"\<lbrakk>r > 0; z \<noteq> 0\<rbrakk> \<Longrightarrow> Ln(z / of_real r) = Ln(z) - ln r"
using Ln_times_of_real [of "inverse r" z]
by (simp add: ln_inverse Ln_of_real mult.commute divide_inverse of_real_inverse [symmetric]
del: of_real_inverse)
-corollary%unimportant Ln_prod:
+corollary\<^marker>\<open>tag unimportant\<close> Ln_prod:
fixes f :: "'a \<Rightarrow> complex"
assumes "finite A" "\<And>x. x \<in> A \<Longrightarrow> f x \<noteq> 0"
shows "\<exists>n. Ln (prod f A) = (\<Sum>x \<in> A. Ln (f x) + (of_int (n x) * (2 * pi)) * \<i>)"
@@ -1753,7 +1753,7 @@
text\<open>Finally: it's is defined for the same interval as the complex logarithm: \<open>(-\<pi>,\<pi>]\<close>.\<close>
-definition%important Arg :: "complex \<Rightarrow> real" where "Arg z \<equiv> (if z = 0 then 0 else Im (Ln z))"
+definition\<^marker>\<open>tag important\<close> Arg :: "complex \<Rightarrow> real" where "Arg z \<equiv> (if z = 0 then 0 else Im (Ln z))"
lemma Arg_of_real: "Arg (of_real r) = (if r<0 then pi else 0)"
by (simp add: Im_Ln_eq_pi Arg_def)
@@ -1891,7 +1891,7 @@
using complex_is_Real_iff
by (simp add: Arg_def Im_Ln_eq_0) (metis less_eq_real_def of_real_Re of_real_def scale_zero_left)
-corollary%unimportant Arg_ne_0: assumes "z \<notin> \<real>\<^sub>\<ge>\<^sub>0" shows "Arg z \<noteq> 0"
+corollary\<^marker>\<open>tag unimportant\<close> Arg_ne_0: assumes "z \<notin> \<real>\<^sub>\<ge>\<^sub>0" shows "Arg z \<noteq> 0"
using assms by (auto simp: nonneg_Reals_def Arg_eq_0)
lemma Arg_eq_pi_iff: "Arg z = pi \<longleftrightarrow> z \<in> \<real> \<and> Re z < 0"
@@ -1993,7 +1993,7 @@
using Arg_exp_diff_2pi [of z]
by (force simp: Ints_def image_def field_simps Arg_def intro!: complex_eqI)
-definition%important unwinding :: "complex \<Rightarrow> int" where
+definition\<^marker>\<open>tag important\<close> unwinding :: "complex \<Rightarrow> int" where
"unwinding z \<equiv> THE k. of_int k = (z - Ln(exp z)) / (of_real(2*pi) * \<i>)"
lemma unwinding: "of_int (unwinding z) = (z - Ln(exp z)) / (of_real(2*pi) * \<i>)"
@@ -2008,7 +2008,7 @@
using unwinding_2pi by (simp add: exp_add)
-subsection%unimportant\<open>Relation between Ln and Arg2pi, and hence continuity of Arg2pi\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Relation between Ln and Arg2pi, and hence continuity of Arg2pi\<close>
lemma Arg2pi_Ln:
assumes "0 < Arg2pi z" shows "Arg2pi z = Im(Ln(-z)) + pi"
@@ -2166,7 +2166,7 @@
using open_Arg2pi2pi_gt [of t]
by (simp add: closed_def Set.Collect_neg_eq [symmetric] not_le)
-subsection%unimportant\<open>Complex Powers\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Complex Powers\<close>
lemma powr_to_1 [simp]: "z powr 1 = (z::complex)"
by (simp add: powr_def)
@@ -2535,7 +2535,7 @@
qed
-subsection%unimportant\<open>Some Limits involving Logarithms\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Some Limits involving Logarithms\<close>
lemma lim_Ln_over_power:
fixes s::complex
@@ -2686,7 +2686,7 @@
qed
-subsection%unimportant\<open>Relation between Square Root and exp/ln, hence its derivative\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Relation between Square Root and exp/ln, hence its derivative\<close>
lemma csqrt_exp_Ln:
assumes "z \<noteq> 0"
@@ -2759,7 +2759,7 @@
"z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> continuous (at z) csqrt"
by (simp add: field_differentiable_within_csqrt field_differentiable_imp_continuous_at)
-corollary%unimportant isCont_csqrt' [simp]:
+corollary\<^marker>\<open>tag unimportant\<close> isCont_csqrt' [simp]:
"\<lbrakk>isCont f z; f z \<notin> \<real>\<^sub>\<le>\<^sub>0\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. csqrt (f x)) z"
by (blast intro: isCont_o2 [OF _ continuous_at_csqrt])
@@ -2802,7 +2802,7 @@
text\<open>The branch cut gives standard bounds in the real case.\<close>
-definition%important Arctan :: "complex \<Rightarrow> complex" where
+definition\<^marker>\<open>tag important\<close> Arctan :: "complex \<Rightarrow> complex" where
"Arctan \<equiv> \<lambda>z. (\<i>/2) * Ln((1 - \<i>*z) / (1 + \<i>*z))"
lemma Arctan_def_moebius: "Arctan z = \<i>/2 * Ln (moebius (-\<i>) 1 \<i> 1 z)"
@@ -3051,7 +3051,7 @@
finally show ?thesis by (subst (asm) sums_of_real_iff)
qed
-subsection%unimportant \<open>Real arctangent\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Real arctangent\<close>
lemma Im_Arctan_of_real [simp]: "Im (Arctan (of_real x)) = 0"
proof -
@@ -3212,7 +3212,7 @@
by (auto simp: arctan_series)
qed
-subsection%unimportant \<open>Bounds on pi using real arctangent\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Bounds on pi using real arctangent\<close>
lemma pi_machin: "pi = 16 * arctan (1 / 5) - 4 * arctan (1 / 239)"
using machin by simp
@@ -3229,7 +3229,7 @@
subsection\<open>Inverse Sine\<close>
-definition%important Arcsin :: "complex \<Rightarrow> complex" where
+definition\<^marker>\<open>tag important\<close> Arcsin :: "complex \<Rightarrow> complex" where
"Arcsin \<equiv> \<lambda>z. -\<i> * Ln(\<i> * z + csqrt(1 - z\<^sup>2))"
lemma Arcsin_body_lemma: "\<i> * z + csqrt(1 - z\<^sup>2) \<noteq> 0"
@@ -3396,7 +3396,7 @@
subsection\<open>Inverse Cosine\<close>
-definition%important Arccos :: "complex \<Rightarrow> complex" where
+definition\<^marker>\<open>tag important\<close> Arccos :: "complex \<Rightarrow> complex" where
"Arccos \<equiv> \<lambda>z. -\<i> * Ln(z + \<i> * csqrt(1 - z\<^sup>2))"
lemma Arccos_range_lemma: "\<bar>Re z\<bar> < 1 \<Longrightarrow> 0 < Im(z + \<i> * csqrt(1 - z\<^sup>2))"
@@ -3555,7 +3555,7 @@
by (simp add: field_differentiable_within_Arccos holomorphic_on_def)
-subsection%unimportant\<open>Upper and Lower Bounds for Inverse Sine and Cosine\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Upper and Lower Bounds for Inverse Sine and Cosine\<close>
lemma Arcsin_bounds: "\<bar>Re z\<bar> < 1 \<Longrightarrow> \<bar>Re(Arcsin z)\<bar> < pi/2"
unfolding Re_Arcsin
@@ -3608,7 +3608,7 @@
qed
-subsection%unimportant\<open>Interrelations between Arcsin and Arccos\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Interrelations between Arcsin and Arccos\<close>
lemma cos_Arcsin_nonzero:
assumes "z\<^sup>2 \<noteq> 1" shows "cos(Arcsin z) \<noteq> 0"
@@ -3710,7 +3710,7 @@
by (simp add: Arcsin_Arccos_csqrt_pos)
-subsection%unimportant\<open>Relationship with Arcsin on the Real Numbers\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Relationship with Arcsin on the Real Numbers\<close>
lemma Im_Arcsin_of_real:
assumes "\<bar>x\<bar> \<le> 1"
@@ -3727,7 +3727,7 @@
by (simp add: Im_Arcsin exp_minus)
qed
-corollary%unimportant Arcsin_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> \<bar>Re z\<bar> \<le> 1 \<Longrightarrow> Arcsin z \<in> \<real>"
+corollary\<^marker>\<open>tag unimportant\<close> Arcsin_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> \<bar>Re z\<bar> \<le> 1 \<Longrightarrow> Arcsin z \<in> \<real>"
by (metis Im_Arcsin_of_real Re_complex_of_real Reals_cases complex_is_Real_iff)
lemma arcsin_eq_Re_Arcsin:
@@ -3760,7 +3760,7 @@
by (metis Im_Arcsin_of_real add.right_neutral arcsin_eq_Re_Arcsin complex_eq mult_zero_right of_real_0)
-subsection%unimportant\<open>Relationship with Arccos on the Real Numbers\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Relationship with Arccos on the Real Numbers\<close>
lemma Im_Arccos_of_real:
assumes "\<bar>x\<bar> \<le> 1"
@@ -3777,7 +3777,7 @@
by (simp add: Im_Arccos exp_minus)
qed
-corollary%unimportant Arccos_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> \<bar>Re z\<bar> \<le> 1 \<Longrightarrow> Arccos z \<in> \<real>"
+corollary\<^marker>\<open>tag unimportant\<close> Arccos_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> \<bar>Re z\<bar> \<le> 1 \<Longrightarrow> Arccos z \<in> \<real>"
by (metis Im_Arccos_of_real Re_complex_of_real Reals_cases complex_is_Real_iff)
lemma arccos_eq_Re_Arccos:
@@ -3809,7 +3809,7 @@
lemma of_real_arccos: "\<bar>x\<bar> \<le> 1 \<Longrightarrow> of_real(arccos x) = Arccos(of_real x)"
by (metis Im_Arccos_of_real add.right_neutral arccos_eq_Re_Arccos complex_eq mult_zero_right of_real_0)
-subsection%unimportant\<open>Some interrelationships among the real inverse trig functions\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Some interrelationships among the real inverse trig functions\<close>
lemma arccos_arctan:
assumes "-1 < x" "x < 1"
@@ -3879,7 +3879,7 @@
using arccos_arcsin_sqrt_pos [of "-x"]
by (simp add: arccos_minus)
-subsection%unimportant\<open>Continuity results for arcsin and arccos\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Continuity results for arcsin and arccos\<close>
lemma continuous_on_Arcsin_real [continuous_intros]:
"continuous_on {w \<in> \<real>. \<bar>Re w\<bar> \<le> 1} Arcsin"
--- a/src/HOL/Analysis/Conformal_Mappings.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Conformal_Mappings.thy Fri Apr 12 22:09:25 2019 +0200
@@ -464,7 +464,7 @@
qed
text\<open>No need for \<^term>\<open>S\<close> to be connected. But the nonconstant condition is stronger.\<close>
-corollary%unimportant open_mapping_thm2:
+corollary\<^marker>\<open>tag unimportant\<close> open_mapping_thm2:
assumes holf: "f holomorphic_on S"
and S: "open S"
and "open U" "U \<subseteq> S"
@@ -494,7 +494,7 @@
by force
qed
-corollary%unimportant open_mapping_thm3:
+corollary\<^marker>\<open>tag unimportant\<close> open_mapping_thm3:
assumes holf: "f holomorphic_on S"
and "open S" and injf: "inj_on f S"
shows "open (f ` S)"
@@ -585,7 +585,7 @@
using z \<open>\<xi> \<in> S\<close> closure_subset by fastforce
qed
-corollary%unimportant maximum_real_frontier:
+corollary\<^marker>\<open>tag unimportant\<close> maximum_real_frontier:
assumes holf: "f holomorphic_on (interior S)"
and contf: "continuous_on (closure S) f"
and bos: "bounded S"
@@ -596,7 +596,7 @@
Transcendental.continuous_on_exp holomorphic_on_compose holomorphic_on_exp assms
by auto
-subsection%unimportant \<open>Factoring out a zero according to its order\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Factoring out a zero according to its order\<close>
lemma holomorphic_factor_order_of_zero:
assumes holf: "f holomorphic_on S"
@@ -1021,7 +1021,7 @@
qed
qed
-subsection%unimportant \<open>Entire proper functions are precisely the non-trivial polynomials\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Entire proper functions are precisely the non-trivial polynomials\<close>
lemma proper_map_polyfun:
fixes c :: "nat \<Rightarrow> 'a::{real_normed_div_algebra,heine_borel}"
@@ -2163,7 +2163,7 @@
text\<open>Wenda Li and LC Paulson (2016). A Formal Proof of Cauchy's Residue Theorem.
Interactive Theorem Proving\<close>
-definition%important residue :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> complex" where
+definition\<^marker>\<open>tag important\<close> residue :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> complex" where
"residue f z = (SOME int. \<exists>e>0. \<forall>\<epsilon>>0. \<epsilon><e
\<longrightarrow> (f has_contour_integral 2*pi* \<i> *int) (circlepath z \<epsilon>))"
@@ -2967,7 +2967,7 @@
subsection \<open>Non-essential singular points\<close>
-definition%important is_pole ::
+definition\<^marker>\<open>tag important\<close> is_pole ::
"('a::topological_space \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> bool" where
"is_pole f a = (LIM x (at a). f x :> at_infinity)"
@@ -3695,12 +3695,12 @@
subsubsection \<open>The order of non-essential singularities (i.e. removable singularities or poles)\<close>
-definition%important zorder :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> int" where
+definition\<^marker>\<open>tag important\<close> zorder :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> int" where
"zorder f z = (THE n. (\<exists>h r. r>0 \<and> h holomorphic_on cball z r \<and> h z\<noteq>0
\<and> (\<forall>w\<in>cball z r - {z}. f w = h w * (w-z) powr (of_int n)
\<and> h w \<noteq>0)))"
-definition%important zor_poly
+definition\<^marker>\<open>tag important\<close> zor_poly
::"[complex \<Rightarrow> complex, complex] \<Rightarrow> complex \<Rightarrow> complex" where
"zor_poly f z = (SOME h. \<exists>r. r > 0 \<and> h holomorphic_on cball z r \<and> h z \<noteq> 0
\<and> (\<forall>w\<in>cball z r - {z}. f w = h w * (w - z) powr (zorder f z)
--- a/src/HOL/Analysis/Connected.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Connected.thy Fri Apr 12 22:09:25 2019 +0200
@@ -9,7 +9,7 @@
Abstract_Topology_2
begin
-subsection%unimportant \<open>Connectedness\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Connectedness\<close>
lemma connected_local:
"connected S \<longleftrightarrow>
@@ -68,7 +68,7 @@
subsection \<open>Connected components, considered as a connectedness relation or a set\<close>
-definition%important "connected_component s x y \<equiv> \<exists>t. connected t \<and> t \<subseteq> s \<and> x \<in> t \<and> y \<in> t"
+definition\<^marker>\<open>tag important\<close> "connected_component s x y \<equiv> \<exists>t. connected t \<and> t \<subseteq> s \<and> x \<in> t \<and> y \<in> t"
abbreviation "connected_component_set s x \<equiv> Collect (connected_component s x)"
@@ -268,7 +268,7 @@
subsection \<open>The set of connected components of a set\<close>
-definition%important components:: "'a::topological_space set \<Rightarrow> 'a set set"
+definition\<^marker>\<open>tag important\<close> components:: "'a::topological_space set \<Rightarrow> 'a set set"
where "components s \<equiv> connected_component_set s ` s"
lemma components_iff: "s \<in> components u \<longleftrightarrow> (\<exists>x. x \<in> u \<and> s = connected_component_set u x)"
@@ -427,7 +427,7 @@
using closedin_connected_component componentsE by blast
-subsection%unimportant \<open>Proving a function is constant on a connected set
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Proving a function is constant on a connected set
by proving that a level set is open\<close>
lemma continuous_levelset_openin_cases:
@@ -458,7 +458,7 @@
by fast
-subsection%unimportant \<open>Preservation of Connectedness\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Preservation of Connectedness\<close>
lemma homeomorphic_connectedness:
assumes "s homeomorphic t"
@@ -688,7 +688,7 @@
qed
-subsection%unimportant\<open>Constancy of a function from a connected set into a finite, disconnected or discrete set\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Constancy of a function from a connected set into a finite, disconnected or discrete set\<close>
text\<open>Still missing: versions for a set that is smaller than R, or countable.\<close>
--- a/src/HOL/Analysis/Continuous_Extension.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Continuous_Extension.thy Fri Apr 12 22:09:25 2019 +0200
@@ -482,7 +482,7 @@
"\<And>x. x \<in> U \<Longrightarrow> norm(g x) \<le> B"
using assms by (auto simp: image_subset_iff intro: Dugundji [of "cball 0 B" U S f])
-corollary%unimportant Tietze_closed_interval:
+corollary\<^marker>\<open>tag unimportant\<close> Tietze_closed_interval:
fixes f :: "'a::{metric_space,second_countable_topology} \<Rightarrow> 'b::euclidean_space"
assumes "continuous_on S f"
and "closedin (top_of_set U) S"
@@ -493,7 +493,7 @@
apply (rule Dugundji [of "cbox a b" U S f])
using assms by auto
-corollary%unimportant Tietze_closed_interval_1:
+corollary\<^marker>\<open>tag unimportant\<close> Tietze_closed_interval_1:
fixes f :: "'a::{metric_space,second_countable_topology} \<Rightarrow> real"
assumes "continuous_on S f"
and "closedin (top_of_set U) S"
@@ -504,7 +504,7 @@
apply (rule Dugundji [of "cbox a b" U S f])
using assms by (auto simp: image_subset_iff)
-corollary%unimportant Tietze_open_interval:
+corollary\<^marker>\<open>tag unimportant\<close> Tietze_open_interval:
fixes f :: "'a::{metric_space,second_countable_topology} \<Rightarrow> 'b::euclidean_space"
assumes "continuous_on S f"
and "closedin (top_of_set U) S"
@@ -515,7 +515,7 @@
apply (rule Dugundji [of "box a b" U S f])
using assms by auto
-corollary%unimportant Tietze_open_interval_1:
+corollary\<^marker>\<open>tag unimportant\<close> Tietze_open_interval_1:
fixes f :: "'a::{metric_space,second_countable_topology} \<Rightarrow> real"
assumes "continuous_on S f"
and "closedin (top_of_set U) S"
@@ -526,7 +526,7 @@
apply (rule Dugundji [of "box a b" U S f])
using assms by (auto simp: image_subset_iff)
-corollary%unimportant Tietze_unbounded:
+corollary\<^marker>\<open>tag unimportant\<close> Tietze_unbounded:
fixes f :: "'a::{metric_space,second_countable_topology} \<Rightarrow> 'b::real_inner"
assumes "continuous_on S f"
and "closedin (top_of_set U) S"
--- a/src/HOL/Analysis/Continuum_Not_Denumerable.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Continuum_Not_Denumerable.thy Fri Apr 12 22:09:25 2019 +0200
@@ -11,7 +11,7 @@
"HOL-Library.Countable_Set"
begin
-subsection%unimportant \<open>Abstract\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Abstract\<close>
text \<open>
The following document presents a proof that the Continuum is uncountable.
@@ -32,7 +32,7 @@
range of \<open>f\<close> by generating a sequence of closed intervals then using the
Nested Interval Property.
\<close>
-text%important \<open>%whitespace\<close>
+text\<^marker>\<open>tag important\<close> \<open>%whitespace\<close>
theorem real_non_denum: "\<nexists>f :: nat \<Rightarrow> real. surj f"
proof
assume "\<exists>f::nat \<Rightarrow> real. surj f"
--- a/src/HOL/Analysis/Convex.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Convex.thy Fri Apr 12 22:09:25 2019 +0200
@@ -16,7 +16,7 @@
subsection \<open>Convexity\<close>
-definition%important convex :: "'a::real_vector set \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> convex :: "'a::real_vector set \<Rightarrow> bool"
where "convex s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u\<ge>0. \<forall>v\<ge>0. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> s)"
lemma convexI:
@@ -189,7 +189,7 @@
by (simp add: convex_def scaleR_conv_of_real)
-subsection%unimportant \<open>Explicit expressions for convexity in terms of arbitrary sums\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Explicit expressions for convexity in terms of arbitrary sums\<close>
lemma convex_sum:
fixes C :: "'a::real_vector set"
@@ -342,7 +342,7 @@
subsection \<open>Functions that are convex on a set\<close>
-definition%important convex_on :: "'a::real_vector set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> convex_on :: "'a::real_vector set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> bool"
where "convex_on s f \<longleftrightarrow>
(\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u\<ge>0. \<forall>v\<ge>0. u + v = 1 \<longrightarrow> f (u *\<^sub>R x + v *\<^sub>R y) \<le> u * f x + v * f y)"
@@ -463,7 +463,7 @@
qed
-subsection%unimportant \<open>Arithmetic operations on sets preserve convexity\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Arithmetic operations on sets preserve convexity\<close>
lemma convex_linear_image:
assumes "linear f"
@@ -929,7 +929,7 @@
qed
-subsection%unimportant \<open>Convexity of real functions\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Convexity of real functions\<close>
lemma convex_on_realI:
assumes "connected A"
@@ -1111,7 +1111,7 @@
subsection \<open>Affine set and affine hull\<close>
-definition%important affine :: "'a::real_vector set \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> affine :: "'a::real_vector set \<Rightarrow> bool"
where "affine s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> s)"
lemma affine_alt: "affine s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u::real. (1 - u) *\<^sub>R x + u *\<^sub>R y \<in> s)"
@@ -1149,7 +1149,7 @@
by (simp add: affine_def algebra_simps) (metis distrib_right mult.left_neutral)
-subsubsection%unimportant \<open>Some explicit formulations\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Some explicit formulations\<close>
text "Formalized by Lars Schewe."
@@ -1305,7 +1305,7 @@
qed
-subsubsection%unimportant \<open>Stepping theorems and hence small special cases\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Stepping theorems and hence small special cases\<close>
lemma affine_hull_empty[simp]: "affine hull {} = {}"
by simp
@@ -1417,7 +1417,7 @@
by (metis add_uminus_conv_diff mem_affine_3_minus real_vector.scale_minus_left)
-subsubsection%unimportant \<open>Some relations between affine hull and subspaces\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Some relations between affine hull and subspaces\<close>
lemma affine_hull_insert_subset_span:
"affine hull (insert a S) \<subseteq> {a + v| v . v \<in> span {x - a | x . x \<in> S}}"
@@ -1473,7 +1473,7 @@
using affine_hull_insert_span[of a "S - {a}", unfolded insert_Diff[OF assms]] by auto
-subsubsection%unimportant \<open>Parallel affine sets\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Parallel affine sets\<close>
definition affine_parallel :: "'a::real_vector set \<Rightarrow> 'a::real_vector set \<Rightarrow> bool"
where "affine_parallel S T \<longleftrightarrow> (\<exists>a. T = (\<lambda>x. a + x) ` S)"
@@ -1573,7 +1573,7 @@
unfolding subspace_def affine_def by auto
-subsubsection%unimportant \<open>Subspace parallel to an affine set\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Subspace parallel to an affine set\<close>
lemma subspace_affine: "subspace S \<longleftrightarrow> affine S \<and> 0 \<in> S"
proof -
@@ -1703,7 +1703,7 @@
subsection \<open>Cones\<close>
-definition%important cone :: "'a::real_vector set \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> cone :: "'a::real_vector set \<Rightarrow> bool"
where "cone s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>c\<ge>0. c *\<^sub>R x \<in> s)"
lemma cone_empty[intro, simp]: "cone {}"
@@ -1869,7 +1869,7 @@
text "Formalized by Lars Schewe."
-definition%important affine_dependent :: "'a::real_vector set \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> affine_dependent :: "'a::real_vector set \<Rightarrow> bool"
where "affine_dependent s \<longleftrightarrow> (\<exists>x\<in>s. x \<in> affine hull (s - {x}))"
lemma affine_dependent_subset:
@@ -1946,7 +1946,7 @@
qed
-subsection%unimportant \<open>Connectedness of convex sets\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Connectedness of convex sets\<close>
lemma connectedD:
"connected S \<Longrightarrow> open A \<Longrightarrow> open B \<Longrightarrow> S \<subseteq> A \<union> B \<Longrightarrow> A \<inter> B \<inter> S = {} \<Longrightarrow> A \<inter> S = {} \<or> B \<inter> S = {}"
@@ -2003,7 +2003,7 @@
lemma convex_hull_eq: "convex hull s = s \<longleftrightarrow> convex s"
by (metis convex_convex_hull hull_same)
-subsubsection%unimportant \<open>Convex hull is "preserved" by a linear function\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Convex hull is "preserved" by a linear function\<close>
lemma convex_hull_linear_image:
assumes f: "linear f"
@@ -2061,7 +2061,7 @@
qed
-subsubsection%unimportant \<open>Stepping theorems for convex hulls of finite sets\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Stepping theorems for convex hulls of finite sets\<close>
lemma convex_hull_empty[simp]: "convex hull {} = {}"
by (rule hull_unique) auto
@@ -2183,7 +2183,7 @@
using diff_eq_eq apply fastforce
by (metis add.group_left_neutral add_le_imp_le_diff diff_add_cancel)
-subsubsection%unimportant \<open>Explicit expression for convex hull\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Explicit expression for convex hull\<close>
proposition convex_hull_indexed:
fixes S :: "'a::real_vector set"
@@ -2262,7 +2262,7 @@
qed (use assms in \<open>auto simp: convex_explicit\<close>)
-subsubsection%unimportant \<open>Another formulation\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Another formulation\<close>
text "Formalized by Lars Schewe."
@@ -2362,7 +2362,7 @@
qed
-subsubsection%unimportant \<open>A stepping theorem for that expansion\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>A stepping theorem for that expansion\<close>
lemma convex_hull_finite_step:
fixes S :: "'a::real_vector set"
@@ -2419,7 +2419,7 @@
qed
-subsubsection%unimportant \<open>Hence some special cases\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Hence some special cases\<close>
lemma convex_hull_2:
"convex hull {a,b} = {u *\<^sub>R a + v *\<^sub>R b | u v. 0 \<le> u \<and> 0 \<le> v \<and> u + v = 1}"
@@ -2491,7 +2491,7 @@
qed
-subsection%unimportant \<open>Relations among closure notions and corresponding hulls\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Relations among closure notions and corresponding hulls\<close>
lemma affine_imp_convex: "affine s \<Longrightarrow> convex s"
unfolding affine_def convex_def by auto
@@ -2629,7 +2629,7 @@
qed
-subsection%unimportant \<open>Some Properties of Affine Dependent Sets\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Some Properties of Affine Dependent Sets\<close>
lemma affine_independent_0 [simp]: "\<not> affine_dependent {}"
by (simp add: affine_dependent_def)
@@ -2868,7 +2868,7 @@
subsection \<open>Affine Dimension of a Set\<close>
-definition%important aff_dim :: "('a::euclidean_space) set \<Rightarrow> int"
+definition\<^marker>\<open>tag important\<close> aff_dim :: "('a::euclidean_space) set \<Rightarrow> int"
where "aff_dim V =
(SOME d :: int.
\<exists>B. affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> of_nat (card B) = d + 1)"
@@ -3656,7 +3656,7 @@
using hull_mono[OF \<open>s\<subseteq>p\<close>] by auto
qed
-subsection%unimportant\<open>Some Properties of subset of standard basis\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Some Properties of subset of standard basis\<close>
lemma affine_hull_substd_basis:
assumes "d \<subseteq> Basis"
@@ -3673,7 +3673,7 @@
by (metis Int_absorb1 Int_absorb2 convex_hull_subset_affine_hull hull_hull hull_mono hull_subset)
-subsection%unimportant \<open>Moving and scaling convex hulls\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Moving and scaling convex hulls\<close>
lemma convex_hull_set_plus:
"convex hull (S + T) = convex hull S + convex hull T"
@@ -3700,7 +3700,7 @@
by(simp only: image_image[symmetric] convex_hull_scaling convex_hull_translation)
-subsection%unimportant \<open>Convexity of cone hulls\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Convexity of cone hulls\<close>
lemma convex_cone_hull:
assumes "convex S"
@@ -4000,7 +4000,7 @@
subsection \<open>Epigraphs of convex functions\<close>
-definition%important "epigraph S (f :: _ \<Rightarrow> real) = {xy. fst xy \<in> S \<and> f (fst xy) \<le> snd xy}"
+definition\<^marker>\<open>tag important\<close> "epigraph S (f :: _ \<Rightarrow> real) = {xy. fst xy \<in> S \<and> f (fst xy) \<le> snd xy}"
lemma mem_epigraph: "(x, y) \<in> epigraph S f \<longleftrightarrow> x \<in> S \<and> f x \<le> y"
unfolding epigraph_def by auto
@@ -4036,7 +4036,7 @@
by (simp add: convex_epigraph)
-subsubsection%unimportant \<open>Use this to derive general bound property of convex function\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Use this to derive general bound property of convex function\<close>
lemma convex_on:
assumes "convex S"
@@ -4060,7 +4060,7 @@
using assms[unfolded convex] apply auto
done
-subsection%unimportant \<open>A bound within a convex hull\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>A bound within a convex hull\<close>
lemma convex_on_convex_hull_bound:
assumes "convex_on (convex hull s) f"
--- a/src/HOL/Analysis/Convex_Euclidean_Space.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy Fri Apr 12 22:09:25 2019 +0200
@@ -14,7 +14,7 @@
Topology_Euclidean_Space
begin
-subsection%unimportant \<open>Topological Properties of Convex Sets and Functions\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Topological Properties of Convex Sets and Functions\<close>
lemma convex_supp_sum:
assumes "convex S" and 1: "supp_sum u I = 1"
@@ -318,7 +318,7 @@
subsection \<open>Relative interior of a set\<close>
-definition%important "rel_interior S =
+definition\<^marker>\<open>tag important\<close> "rel_interior S =
{x. \<exists>T. openin (top_of_set (affine hull S)) T \<and> x \<in> T \<and> T \<subseteq> S}"
lemma rel_interior_mono:
@@ -649,7 +649,7 @@
subsubsection \<open>Relative open sets\<close>
-definition%important "rel_open S \<longleftrightarrow> rel_interior S = S"
+definition\<^marker>\<open>tag important\<close> "rel_open S \<longleftrightarrow> rel_interior S = S"
lemma rel_open: "rel_open S \<longleftrightarrow> openin (top_of_set (affine hull S)) S"
unfolding rel_open_def rel_interior_def
@@ -848,7 +848,7 @@
qed
-subsubsection%unimportant\<open>Relative interior preserves under linear transformations\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close>\<open>Relative interior preserves under linear transformations\<close>
lemma rel_interior_translation_aux:
fixes a :: "'n::euclidean_space"
@@ -1005,7 +1005,7 @@
by auto
-subsection%unimportant \<open>Openness and compactness are preserved by convex hull operation\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Openness and compactness are preserved by convex hull operation\<close>
lemma open_convex_hull[intro]:
fixes S :: "'a::real_normed_vector set"
@@ -1237,7 +1237,7 @@
qed
-subsection%unimportant \<open>Extremal points of a simplex are some vertices\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Extremal points of a simplex are some vertices\<close>
lemma dist_increases_online:
fixes a b d :: "'a::real_inner"
@@ -1432,7 +1432,7 @@
subsection \<open>Closest point of a convex set is unique, with a continuous projection\<close>
-definition%important closest_point :: "'a::{real_inner,heine_borel} set \<Rightarrow> 'a \<Rightarrow> 'a"
+definition\<^marker>\<open>tag important\<close> closest_point :: "'a::{real_inner,heine_borel} set \<Rightarrow> 'a \<Rightarrow> 'a"
where "closest_point S a = (SOME x. x \<in> S \<and> (\<forall>y\<in>S. dist a x \<le> dist a y))"
lemma closest_point_exists:
@@ -1622,7 +1622,7 @@
qed
-subsubsection%unimportant \<open>Various point-to-set separating/supporting hyperplane theorems\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Various point-to-set separating/supporting hyperplane theorems\<close>
lemma supporting_hyperplane_closed_point:
fixes z :: "'a::{real_inner,heine_borel}"
@@ -1707,7 +1707,7 @@
qed
-subsubsection%unimportant \<open>Now set-to-set for closed/compact sets\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Now set-to-set for closed/compact sets\<close>
lemma separating_hyperplane_closed_compact:
fixes S :: "'a::euclidean_space set"
@@ -1797,7 +1797,7 @@
qed
-subsubsection%unimportant \<open>General case without assuming closure and getting non-strict separation\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>General case without assuming closure and getting non-strict separation\<close>
lemma separating_hyperplane_set_0:
assumes "convex S" "(0::'a::euclidean_space) \<notin> S"
@@ -1855,7 +1855,7 @@
qed
-subsection%unimportant \<open>More convexity generalities\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>More convexity generalities\<close>
lemma convex_closure [intro,simp]:
fixes S :: "'a::real_normed_vector set"
@@ -1898,7 +1898,7 @@
using hull_subset[of S convex] convex_hull_empty by auto
-subsection%unimportant \<open>Convex set as intersection of halfspaces\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Convex set as intersection of halfspaces\<close>
lemma convex_halfspace_intersection:
fixes s :: "('a::euclidean_space) set"
@@ -1922,7 +1922,7 @@
qed auto
-subsection%unimportant \<open>Convexity of general and special intervals\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Convexity of general and special intervals\<close>
lemma is_interval_convex:
fixes S :: "'a::euclidean_space set"
@@ -2007,7 +2007,7 @@
"\<lbrakk>connected S; aff_dim S \<noteq> 0; a \<in> S\<rbrakk> \<Longrightarrow> a islimpt S"
using aff_dim_sing connected_imp_perfect by blast
-subsection%unimportant \<open>On \<open>real\<close>, \<open>is_interval\<close>, \<open>convex\<close> and \<open>connected\<close> are all equivalent\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>On \<open>real\<close>, \<open>is_interval\<close>, \<open>convex\<close> and \<open>connected\<close> are all equivalent\<close>
lemma mem_is_interval_1_I:
fixes a b c::real
@@ -2099,7 +2099,7 @@
and is_interval_co: "is_interval {r..<s}"
by (simp_all add: is_interval_convex_1)
-subsection%unimportant \<open>Another intermediate value theorem formulation\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Another intermediate value theorem formulation\<close>
lemma ivt_increasing_component_on_1:
fixes f :: "real \<Rightarrow> 'a::euclidean_space"
@@ -2144,7 +2144,7 @@
by (rule ivt_decreasing_component_on_1) (auto simp: continuous_at_imp_continuous_on)
-subsection%unimportant \<open>A bound within an interval\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>A bound within an interval\<close>
lemma convex_hull_eq_real_cbox:
fixes x y :: real assumes "x \<le> y"
@@ -2235,7 +2235,7 @@
by (rule_tac that[of "(\<lambda>y. x - ?d + (2 * d) *\<^sub>R y)` S"]) (auto simp: convex_hull_affinity *)
qed
-subsection%unimportant\<open>Representation of any interval as a finite convex hull\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Representation of any interval as a finite convex hull\<close>
lemma image_stretch_interval:
"(\<lambda>x. \<Sum>k\<in>Basis. (m k * (x\<bullet>k)) *\<^sub>R k) ` cbox a (b::'a::euclidean_space) =
@@ -2312,7 +2312,7 @@
qed
-subsection%unimportant \<open>Bounded convex function on open set is continuous\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Bounded convex function on open set is continuous\<close>
lemma convex_on_bounded_continuous:
fixes S :: "('a::real_normed_vector) set"
@@ -2405,7 +2405,7 @@
qed
-subsection%unimportant \<open>Upper bound on a ball implies upper and lower bounds\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Upper bound on a ball implies upper and lower bounds\<close>
lemma convex_bounds_lemma:
fixes x :: "'a::real_normed_vector"
@@ -2439,7 +2439,7 @@
qed
-subsubsection%unimportant \<open>Hence a convex function on an open set is continuous\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Hence a convex function on an open set is continuous\<close>
lemma real_of_nat_ge_one_iff: "1 \<le> real (n::nat) \<longleftrightarrow> 1 \<le> n"
by auto
--- a/src/HOL/Analysis/Cross3.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Cross3.thy Fri Apr 12 22:09:25 2019 +0200
@@ -13,7 +13,7 @@
context includes no_Set_Product_syntax
begin \<comment>\<open>locally disable syntax for set product, to avoid warnings\<close>
-definition%important cross3 :: "[real^3, real^3] \<Rightarrow> real^3" (infixr "\<times>" 80)
+definition\<^marker>\<open>tag important\<close> cross3 :: "[real^3, real^3] \<Rightarrow> real^3" (infixr "\<times>" 80)
where "a \<times> b \<equiv>
vector [a$2 * b$3 - a$3 * b$2,
a$3 * b$1 - a$1 * b$3,
--- a/src/HOL/Analysis/Derivative.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Derivative.thy Fri Apr 12 22:09:25 2019 +0200
@@ -20,7 +20,7 @@
by (intro derivative_eq_intros) auto
-subsection%unimportant \<open>Derivative with composed bilinear function\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Derivative with composed bilinear function\<close>
text \<open>More explicit epsilon-delta forms.\<close>
@@ -92,7 +92,7 @@
subsection \<open>Differentiability\<close>
-definition%important
+definition\<^marker>\<open>tag important\<close>
differentiable_on :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a set \<Rightarrow> bool"
(infix "differentiable'_on" 50)
where "f differentiable_on s \<longleftrightarrow> (\<forall>x\<in>s. f differentiable (at x within s))"
@@ -113,7 +113,7 @@
"(\<And>x. x \<in> s \<Longrightarrow> f differentiable at x) \<Longrightarrow> f differentiable_on s"
by (metis differentiable_at_withinI differentiable_on_def)
-corollary%unimportant differentiable_iff_scaleR:
+corollary\<^marker>\<open>tag unimportant\<close> differentiable_iff_scaleR:
fixes f :: "real \<Rightarrow> 'a::real_normed_vector"
shows "f differentiable F \<longleftrightarrow> (\<exists>d. (f has_derivative (\<lambda>x. x *\<^sub>R d)) F)"
by (auto simp: differentiable_def dest: has_derivative_linear linear_imp_scaleR)
@@ -324,7 +324,7 @@
by (auto simp: o_def mult.commute has_vector_derivative_def)
-subsection%unimportant \<open>Composition rules stated just for differentiability\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Composition rules stated just for differentiability\<close>
lemma differentiable_chain_at:
"f differentiable (at x) \<Longrightarrow>
@@ -342,7 +342,7 @@
subsection \<open>Uniqueness of derivative\<close>
-text%important \<open>
+text\<^marker>\<open>tag important\<close> \<open>
The general result is a bit messy because we need approachability of the
limit point from any direction. But OK for nontrivial intervals etc.
\<close>
@@ -2195,7 +2195,7 @@
subsection \<open>Field differentiability\<close>
-definition%important field_differentiable :: "['a \<Rightarrow> 'a::real_normed_field, 'a filter] \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> field_differentiable :: "['a \<Rightarrow> 'a::real_normed_field, 'a filter] \<Rightarrow> bool"
(infixr "(field'_differentiable)" 50)
where "f field_differentiable F \<equiv> \<exists>f'. (f has_field_derivative f') F"
@@ -2390,7 +2390,7 @@
subsection \<open>Field derivative\<close>
-definition%important deriv :: "('a \<Rightarrow> 'a::real_normed_field) \<Rightarrow> 'a \<Rightarrow> 'a" where
+definition\<^marker>\<open>tag important\<close> deriv :: "('a \<Rightarrow> 'a::real_normed_field) \<Rightarrow> 'a \<Rightarrow> 'a" where
"deriv f x \<equiv> SOME D. DERIV f x :> D"
lemma DERIV_imp_deriv: "DERIV f x :> f' \<Longrightarrow> deriv f x = f'"
@@ -2701,7 +2701,7 @@
qed
-subsection%unimportant \<open>Differentiable case distinction\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Differentiable case distinction\<close>
lemma has_derivative_within_If_eq:
"((\<lambda>x. if P x then f x else g x) has_derivative f') (at x within S) =
--- a/src/HOL/Analysis/Determinants.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Determinants.thy Fri Apr 12 22:09:25 2019 +0200
@@ -12,7 +12,7 @@
subsection \<open>Trace\<close>
-definition%important trace :: "'a::semiring_1^'n^'n \<Rightarrow> 'a"
+definition\<^marker>\<open>tag important\<close> trace :: "'a::semiring_1^'n^'n \<Rightarrow> 'a"
where "trace A = sum (\<lambda>i. ((A$i)$i)) (UNIV::'n set)"
lemma trace_0: "trace (mat 0) = 0"
@@ -33,9 +33,9 @@
apply (simp add: mult.commute)
done
-subsubsection%important \<open>Definition of determinant\<close>
+subsubsection\<^marker>\<open>tag important\<close> \<open>Definition of determinant\<close>
-definition%important det:: "'a::comm_ring_1^'n^'n \<Rightarrow> 'a" where
+definition\<^marker>\<open>tag important\<close> det:: "'a::comm_ring_1^'n^'n \<Rightarrow> 'a" where
"det A =
sum (\<lambda>p. of_int (sign p) * prod (\<lambda>i. A$i$p i) (UNIV :: 'n set))
{p. p permutes (UNIV :: 'n set)}"
@@ -776,7 +776,7 @@
using invertible_det_nz [of A]
by (auto simp: matrix_left_invertible_injective invertible_left_inverse less_rank_noninjective)
-subsubsection%important \<open>Invertibility of matrices and corresponding linear functions\<close>
+subsubsection\<^marker>\<open>tag important\<close> \<open>Invertibility of matrices and corresponding linear functions\<close>
lemma matrix_left_invertible_gen:
fixes f :: "'a::field^'m \<Rightarrow> 'a::field^'n"
@@ -994,12 +994,12 @@
proposition orthogonal_transformation_det [simp]:
fixes f :: "real^'n \<Rightarrow> real^'n"
shows "orthogonal_transformation f \<Longrightarrow> \<bar>det (matrix f)\<bar> = 1"
- using%unimportant det_orthogonal_matrix orthogonal_transformation_matrix by fastforce
+ using det_orthogonal_matrix orthogonal_transformation_matrix by fastforce
subsection \<open>Rotation, reflection, rotoinversion\<close>
-definition%important "rotation_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = 1"
-definition%important "rotoinversion_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = - 1"
+definition\<^marker>\<open>tag important\<close> "rotation_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = 1"
+definition\<^marker>\<open>tag important\<close> "rotoinversion_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = - 1"
lemma orthogonal_rotation_or_rotoinversion:
fixes Q :: "'a::linordered_idom^'n^'n"
--- a/src/HOL/Analysis/Elementary_Metric_Spaces.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Elementary_Metric_Spaces.thy Fri Apr 12 22:09:25 2019 +0200
@@ -15,13 +15,13 @@
subsection \<open>Open and closed balls\<close>
-definition%important ball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set"
+definition\<^marker>\<open>tag important\<close> ball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set"
where "ball x e = {y. dist x y < e}"
-definition%important cball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set"
+definition\<^marker>\<open>tag important\<close> cball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set"
where "cball x e = {y. dist x y \<le> e}"
-definition%important sphere :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set"
+definition\<^marker>\<open>tag important\<close> sphere :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set"
where "sphere x e = {y. dist x y = e}"
lemma mem_ball [simp]: "y \<in> ball x e \<longleftrightarrow> dist x y < e"
@@ -717,7 +717,7 @@
subsection \<open>Boundedness\<close>
(* FIXME: This has to be unified with BSEQ!! *)
-definition%important (in metric_space) bounded :: "'a set \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> (in metric_space) bounded :: "'a set \<Rightarrow> bool"
where "bounded S \<longleftrightarrow> (\<exists>x e. \<forall>y\<in>S. dist x y \<le> e)"
lemma bounded_subset_cball: "bounded S \<longleftrightarrow> (\<exists>e x. S \<subseteq> cball x e \<and> 0 \<le> e)"
@@ -1198,7 +1198,7 @@
subsection \<open>The diameter of a set\<close>
-definition%important diameter :: "'a::metric_space set \<Rightarrow> real" where
+definition\<^marker>\<open>tag important\<close> diameter :: "'a::metric_space set \<Rightarrow> real" where
"diameter S = (if S = {} then 0 else SUP (x,y)\<in>S\<times>S. dist x y)"
lemma diameter_empty [simp]: "diameter{} = 0"
@@ -1464,8 +1464,8 @@
shows "compact(closure S) \<longleftrightarrow> bounded S"
by (meson bounded_closure bounded_subset closed_closure closure_subset compact_eq_bounded_closed)
-instance%important real :: heine_borel
-proof%unimportant
+instance\<^marker>\<open>tag important\<close> real :: heine_borel
+proof
fix f :: "nat \<Rightarrow> real"
assume f: "bounded (range f)"
obtain r :: "nat \<Rightarrow> nat" where r: "strict_mono r" "monoseq (f \<circ> r)"
@@ -1545,8 +1545,8 @@
unfolding bounded_def
by (metis (no_types, hide_lams) dist_snd_le image_iff order.trans)
-instance%important prod :: (heine_borel, heine_borel) heine_borel
-proof%unimportant
+instance\<^marker>\<open>tag important\<close> prod :: (heine_borel, heine_borel) heine_borel
+proof
fix f :: "nat \<Rightarrow> 'a \<times> 'b"
assume f: "bounded (range f)"
then have "bounded (fst ` range f)"
@@ -1844,7 +1844,7 @@
by auto
-subsection%unimportant\<open> Finite intersection property\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open> Finite intersection property\<close>
text\<open>Also developed in HOL's toplogical spaces theory, but the Heine-Borel type class isn't available there.\<close>
@@ -1961,7 +1961,7 @@
subsection \<open>Infimum Distance\<close>
-definition%important "infdist x A = (if A = {} then 0 else INF a\<in>A. dist x a)"
+definition\<^marker>\<open>tag important\<close> "infdist x A = (if A = {} then 0 else INF a\<in>A. dist x a)"
lemma bdd_below_image_dist[intro, simp]: "bdd_below (dist x ` A)"
by (auto intro!: zero_le_dist)
@@ -2453,7 +2453,7 @@
by (force intro: compact_uniformly_equicontinuous [OF S, of "{f}"])
-subsection%unimportant\<open> Theorems relating continuity and uniform continuity to closures\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open> Theorems relating continuity and uniform continuity to closures\<close>
lemma continuous_on_closure:
"continuous_on (closure S) f \<longleftrightarrow>
@@ -2860,7 +2860,7 @@
then show ?thesis ..
qed
-subsection%unimportant \<open>Making a continuous function avoid some value in a neighbourhood\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Making a continuous function avoid some value in a neighbourhood\<close>
lemma continuous_within_avoid:
fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space"
@@ -3041,7 +3041,7 @@
subsection\<open>The infimum of the distance between two sets\<close>
-definition%important setdist :: "'a::metric_space set \<Rightarrow> 'a set \<Rightarrow> real" where
+definition\<^marker>\<open>tag important\<close> setdist :: "'a::metric_space set \<Rightarrow> 'a set \<Rightarrow> real" where
"setdist s t \<equiv>
(if s = {} \<or> t = {} then 0
else Inf {dist x y| x y. x \<in> s \<and> y \<in> t})"
--- a/src/HOL/Analysis/Elementary_Normed_Spaces.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Elementary_Normed_Spaces.thy Fri Apr 12 22:09:25 2019 +0200
@@ -13,7 +13,7 @@
Connected
begin
-subsection%unimportant \<open>Various Lemmas Combining Imports\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Various Lemmas Combining Imports\<close>
lemma countable_PiE:
"finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> countable (F i)) \<Longrightarrow> countable (Pi\<^sub>E I F)"
@@ -428,7 +428,7 @@
done
-subsection%unimportant \<open>Various Lemmas on Normed Algebras\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Various Lemmas on Normed Algebras\<close>
lemma closed_of_nat_image: "closed (of_nat ` A :: 'a::real_normed_algebra_1 set)"
@@ -606,7 +606,7 @@
using LIM_offset_zero LIM_offset_zero_cancel ..
-subsection%unimportant \<open>Limit Point of Filter\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Limit Point of Filter\<close>
lemma netlimit_at_vector:
fixes a :: "'a::real_normed_vector"
@@ -775,7 +775,7 @@
shows "bounded (- S) \<Longrightarrow> \<not> bounded S"
using bounded_Un [of S "-S"] by (simp add: sup_compl_top)
-subsection%unimportant\<open>Relations among convergence and absolute convergence for power series\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Relations among convergence and absolute convergence for power series\<close>
lemma summable_imp_bounded:
fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
@@ -1024,7 +1024,7 @@
subsection \<open>Continuity\<close>
-subsubsection%unimportant \<open>Structural rules for uniform continuity\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Structural rules for uniform continuity\<close>
lemma (in bounded_linear) uniformly_continuous_on[continuous_intros]:
fixes g :: "_::metric_space \<Rightarrow> _"
@@ -1103,7 +1103,7 @@
by (simp add: fun_Compl_def uniformly_continuous_on_minus)
-subsection%unimportant \<open>Topological properties of linear functions\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Topological properties of linear functions\<close>
lemma linear_lim_0:
assumes "bounded_linear f"
@@ -1131,7 +1131,7 @@
"bounded_linear f \<Longrightarrow> continuous_on s f"
using continuous_at_imp_continuous_on[of s f] using linear_continuous_at[of f] by auto
-subsection%unimportant \<open>Arithmetic Preserves Topological Properties\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Arithmetic Preserves Topological Properties\<close>
lemma open_scaling[intro]:
fixes s :: "'a::real_normed_vector set"
@@ -1487,7 +1487,7 @@
using ball_translation [of "- a" c] by (simp cong: image_cong_simp)
-subsection%unimportant\<open>Homeomorphisms\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Homeomorphisms\<close>
lemma homeomorphic_scaling:
fixes s :: "'a::real_normed_vector set"
@@ -1597,7 +1597,7 @@
using assms homeomorphic_ball01_UNIV homeomorphic_balls(1) homeomorphic_trans zero_less_one by blast
-subsection%unimportant \<open>Discrete\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Discrete\<close>
lemma finite_implies_discrete:
fixes S :: "'a::topological_space set"
@@ -1627,7 +1627,7 @@
qed
-subsection%unimportant \<open>Completeness of "Isometry" (up to constant bounds)\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Completeness of "Isometry" (up to constant bounds)\<close>
lemma cauchy_isometric:\<comment> \<open>TODO: rename lemma to \<open>Cauchy_isometric\<close>\<close>
assumes e: "e > 0"
--- a/src/HOL/Analysis/Elementary_Topology.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Elementary_Topology.thy Fri Apr 12 22:09:25 2019 +0200
@@ -22,7 +22,7 @@
using openI by auto
-subsubsection%unimportant \<open>Archimedean properties and useful consequences\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Archimedean properties and useful consequences\<close>
text\<open>Bernoulli's inequality\<close>
proposition Bernoulli_inequality:
@@ -111,7 +111,7 @@
apply (metis Suc_pred of_nat_Suc)
done
-subsubsection%unimportant \<open>Affine transformations of intervals\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Affine transformations of intervals\<close>
lemma real_affinity_le: "0 < m \<Longrightarrow> m * x + c \<le> y \<longleftrightarrow> x \<le> inverse m * y + - (c / m)"
for m :: "'a::linordered_field"
@@ -144,7 +144,7 @@
context topological_space
begin
-definition%important "topological_basis B \<longleftrightarrow>
+definition\<^marker>\<open>tag important\<close> "topological_basis B \<longleftrightarrow>
(\<forall>b\<in>B. open b) \<and> (\<forall>x. open x \<longrightarrow> (\<exists>B'. B' \<subseteq> B \<and> \<Union>B' = x))"
lemma topological_basis:
@@ -270,7 +270,7 @@
subsection \<open>Countable Basis\<close>
-locale%important countable_basis = topological_space p for p::"'a set \<Rightarrow> bool" +
+locale\<^marker>\<open>tag important\<close> countable_basis = topological_space p for p::"'a set \<Rightarrow> bool" +
fixes B :: "'a set set"
assumes is_basis: "topological_basis B"
and countable_basis: "countable B"
@@ -666,7 +666,7 @@
subsection \<open>Limit Points\<close>
-definition%important (in topological_space) islimpt:: "'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixr "islimpt" 60)
+definition\<^marker>\<open>tag important\<close> (in topological_space) islimpt:: "'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixr "islimpt" 60)
where "x islimpt S \<longleftrightarrow> (\<forall>T. x\<in>T \<longrightarrow> open T \<longrightarrow> (\<exists>y\<in>S. y\<in>T \<and> y\<noteq>x))"
lemma islimptI:
@@ -866,7 +866,7 @@
subsection \<open>Interior of a Set\<close>
-definition%important interior :: "('a::topological_space) set \<Rightarrow> 'a set" where
+definition\<^marker>\<open>tag important\<close> interior :: "('a::topological_space) set \<Rightarrow> 'a set" where
"interior S = \<Union>{T. open T \<and> T \<subseteq> S}"
lemma interiorI [intro?]:
@@ -1045,7 +1045,7 @@
subsection \<open>Closure of a Set\<close>
-definition%important closure :: "('a::topological_space) set \<Rightarrow> 'a set" where
+definition\<^marker>\<open>tag important\<close> closure :: "('a::topological_space) set \<Rightarrow> 'a set" where
"closure S = S \<union> {x . x islimpt S}"
lemma interior_closure: "interior S = - (closure (- S))"
@@ -1175,7 +1175,7 @@
subsection \<open>Frontier (also known as boundary)\<close>
-definition%important frontier :: "('a::topological_space) set \<Rightarrow> 'a set" where
+definition\<^marker>\<open>tag important\<close> frontier :: "('a::topological_space) set \<Rightarrow> 'a set" where
"frontier S = closure S - interior S"
lemma frontier_closed [iff]: "closed (frontier S)"
@@ -1253,7 +1253,7 @@
qed
-subsection%unimportant \<open>Filters and the ``eventually true'' quantifier\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Filters and the ``eventually true'' quantifier\<close>
text \<open>Identify Trivial limits, where we can't approach arbitrarily closely.\<close>
@@ -1846,7 +1846,7 @@
with \<open>U \<inter> \<Inter>A = {}\<close> show False by auto
qed
-definition%important countably_compact :: "('a::topological_space) set \<Rightarrow> bool" where
+definition\<^marker>\<open>tag important\<close> countably_compact :: "('a::topological_space) set \<Rightarrow> bool" where
"countably_compact U \<longleftrightarrow>
(\<forall>A. countable A \<longrightarrow> (\<forall>a\<in>A. open a) \<longrightarrow> U \<subseteq> \<Union>A
\<longrightarrow> (\<exists>T\<subseteq>A. finite T \<and> U \<subseteq> \<Union>T))"
@@ -1916,7 +1916,7 @@
subsubsection\<open>Sequential compactness\<close>
-definition%important seq_compact :: "'a::topological_space set \<Rightarrow> bool" where
+definition\<^marker>\<open>tag important\<close> seq_compact :: "'a::topological_space set \<Rightarrow> bool" where
"seq_compact S \<longleftrightarrow>
(\<forall>f. (\<forall>n. f n \<in> S)
\<longrightarrow> (\<exists>l\<in>S. \<exists>r::nat\<Rightarrow>nat. strict_mono r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially))"
@@ -2184,7 +2184,7 @@
by (rule countable_acc_point_imp_seq_compact) (metis islimpt_eq_acc_point)
-subsection%unimportant \<open>Cartesian products\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Cartesian products\<close>
lemma seq_compact_Times: "seq_compact s \<Longrightarrow> seq_compact t \<Longrightarrow> seq_compact (s \<times> t)"
unfolding seq_compact_def
@@ -2457,7 +2457,7 @@
subsection \<open>Homeomorphisms\<close>
-definition%important "homeomorphism s t f g \<longleftrightarrow>
+definition\<^marker>\<open>tag important\<close> "homeomorphism s t f g \<longleftrightarrow>
(\<forall>x\<in>s. (g(f x) = x)) \<and> (f ` s = t) \<and> continuous_on s f \<and>
(\<forall>y\<in>t. (f(g y) = y)) \<and> (g ` t = s) \<and> continuous_on t g"
@@ -2489,7 +2489,7 @@
lemma homeomorphism_sym: "homeomorphism S t f g = homeomorphism t S g f"
by (force simp: homeomorphism_def)
-definition%important homeomorphic :: "'a::topological_space set \<Rightarrow> 'b::topological_space set \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> homeomorphic :: "'a::topological_space set \<Rightarrow> 'b::topological_space set \<Rightarrow> bool"
(infixr "homeomorphic" 60)
where "s homeomorphic t \<equiv> (\<exists>f g. homeomorphism s t f g)"
@@ -2661,7 +2661,7 @@
by (metis compact_continuous_image)
-subsection%unimportant \<open>On Linorder Topologies\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>On Linorder Topologies\<close>
lemma islimpt_greaterThanLessThan1:
fixes a b::"'a::{linorder_topology, dense_order}"
--- a/src/HOL/Analysis/Embed_Measure.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Embed_Measure.thy Fri Apr 12 22:09:25 2019 +0200
@@ -22,7 +22,7 @@
including a constructor \<open>RealVal\<close> for real numbers. Then \<open>embed_measure\<close> allows us to lift a
measure on real numbers to the appropriate subset of that algebraic datatype.
\<close>
-definition%important embed_measure :: "'a measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b measure" where
+definition\<^marker>\<open>tag important\<close> embed_measure :: "'a measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b measure" where
"embed_measure M f = measure_of (f ` space M) {f ` A |A. A \<in> sets M}
(\<lambda>A. emeasure M (f -` A \<inter> space M))"
--- a/src/HOL/Analysis/Euclidean_Space.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Euclidean_Space.thy Fri Apr 12 22:09:25 2019 +0200
@@ -13,7 +13,7 @@
begin
-subsection%unimportant \<open>Interlude: Some properties of real sets\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Interlude: Some properties of real sets\<close>
lemma seq_mono_lemma:
assumes "\<forall>(n::nat) \<ge> m. (d n :: real) < e n"
@@ -211,7 +211,7 @@
qed
-subsection%unimportant \<open>Subclass relationships\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Subclass relationships\<close>
instance euclidean_space \<subseteq> perfect_space
proof
@@ -234,7 +234,7 @@
subsection \<open>Class instances\<close>
-subsubsection%unimportant \<open>Type \<^typ>\<open>real\<close>\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Type \<^typ>\<open>real\<close>\<close>
instantiation real :: euclidean_space
begin
@@ -250,7 +250,7 @@
lemma DIM_real[simp]: "DIM(real) = 1"
by simp
-subsubsection%unimportant \<open>Type \<^typ>\<open>complex\<close>\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Type \<^typ>\<open>complex\<close>\<close>
instantiation complex :: euclidean_space
begin
@@ -271,7 +271,7 @@
lemma complex_Basis_i [iff]: "\<i> \<in> Basis"
by (simp add: Basis_complex_def)
-subsubsection%unimportant \<open>Type \<^typ>\<open>'a \<times> 'b\<close>\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Type \<^typ>\<open>'a \<times> 'b\<close>\<close>
instantiation prod :: (real_inner, real_inner) real_inner
begin
--- a/src/HOL/Analysis/Extended_Real_Limits.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Extended_Real_Limits.thy Fri Apr 12 22:09:25 2019 +0200
@@ -24,7 +24,7 @@
lemma compact_eq_closed:
fixes S :: "'a::{complete_linorder,linorder_topology,second_countable_topology} set"
shows "compact S \<longleftrightarrow> closed S"
- using%unimportant closed_Int_compact[of S, OF _ compact_UNIV] compact_imp_closed
+ using closed_Int_compact[of S, OF _ compact_UNIV] compact_imp_closed
by auto
lemma closed_contains_Sup_cl:
@@ -57,7 +57,7 @@
by simp
qed
-instance%unimportant enat :: second_countable_topology
+instance\<^marker>\<open>tag unimportant\<close> enat :: second_countable_topology
proof
show "\<exists>B::enat set set. countable B \<and> open = generate_topology B"
proof (intro exI conjI)
@@ -66,7 +66,7 @@
qed (simp add: open_enat_def)
qed
-instance%unimportant ereal :: second_countable_topology
+instance\<^marker>\<open>tag unimportant\<close> ereal :: second_countable_topology
proof (standard, intro exI conjI)
let ?B = "(\<Union>r\<in>\<rat>. {{..< r}, {r <..}} :: ereal set set)"
show "countable ?B"
@@ -391,7 +391,7 @@
qed
-subsubsection%important \<open>Continuity of addition\<close>
+subsubsection\<^marker>\<open>tag important\<close> \<open>Continuity of addition\<close>
text \<open>The next few lemmas remove an unnecessary assumption in \<open>tendsto_add_ereal\<close>, culminating
in \<open>tendsto_add_ereal_general\<close> which essentially says that the addition
@@ -522,7 +522,7 @@
then show ?thesis using tendsto_add_ereal_MInf MInf f g by (metis ereal_MInfty_eq_plus)
qed
-subsubsection%important \<open>Continuity of multiplication\<close>
+subsubsection\<^marker>\<open>tag important\<close> \<open>Continuity of multiplication\<close>
text \<open>In the same way as for addition, we prove that the multiplication is continuous on
ereal times ereal, except at \<open>(\<infinity>, 0)\<close> and \<open>(-\<infinity>, 0)\<close> and \<open>(0, \<infinity>)\<close> and \<open>(0, -\<infinity>)\<close>,
@@ -659,7 +659,7 @@
by (cases "c = 0", auto simp add: assms tendsto_mult_ereal)
-subsubsection%important \<open>Continuity of division\<close>
+subsubsection\<^marker>\<open>tag important\<close> \<open>Continuity of division\<close>
lemma tendsto_inverse_ereal_PInf:
fixes u::"_ \<Rightarrow> ereal"
--- a/src/HOL/Analysis/FPS_Convergence.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/FPS_Convergence.thy Fri Apr 12 22:09:25 2019 +0200
@@ -14,7 +14,7 @@
"HOL-Computational_Algebra.Formal_Power_Series"
begin
-subsection%unimportant \<open>Balls with extended real radius\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Balls with extended real radius\<close>
text \<open>
The following is a variant of \<^const>\<open>ball\<close> that also allows an infinite radius.
@@ -55,13 +55,13 @@
subsection \<open>Basic properties of convergent power series\<close>
-definition%important fps_conv_radius :: "'a :: {banach, real_normed_div_algebra} fps \<Rightarrow> ereal" where
+definition\<^marker>\<open>tag important\<close> fps_conv_radius :: "'a :: {banach, real_normed_div_algebra} fps \<Rightarrow> ereal" where
"fps_conv_radius f = conv_radius (fps_nth f)"
-definition%important eval_fps :: "'a :: {banach, real_normed_div_algebra} fps \<Rightarrow> 'a \<Rightarrow> 'a" where
+definition\<^marker>\<open>tag important\<close> eval_fps :: "'a :: {banach, real_normed_div_algebra} fps \<Rightarrow> 'a \<Rightarrow> 'a" where
"eval_fps f z = (\<Sum>n. fps_nth f n * z ^ n)"
-definition%important fps_expansion :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> complex fps" where
+definition\<^marker>\<open>tag important\<close> fps_expansion :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> complex fps" where
"fps_expansion f z0 = Abs_fps (\<lambda>n. (deriv ^^ n) f z0 / fact n)"
lemma norm_summable_fps:
@@ -213,7 +213,7 @@
qed
-subsection%unimportant \<open>Lower bounds on radius of convergence\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Lower bounds on radius of convergence\<close>
lemma fps_conv_radius_deriv:
fixes f :: "'a :: {banach, real_normed_field} fps"
@@ -811,7 +811,7 @@
the coefficients of the series with the singularities of the function, this predicate
is enough.
\<close>
-definition%important
+definition\<^marker>\<open>tag important\<close>
has_fps_expansion :: "('a :: {banach,real_normed_div_algebra} \<Rightarrow> 'a) \<Rightarrow> 'a fps \<Rightarrow> bool"
(infixl "has'_fps'_expansion" 60)
where "(f has_fps_expansion F) \<longleftrightarrow>
--- a/src/HOL/Analysis/Fashoda_Theorem.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Fashoda_Theorem.thy Fri Apr 12 22:09:25 2019 +0200
@@ -10,7 +10,7 @@
subsection \<open>Bijections between intervals\<close>
-definition%important interval_bij :: "'a \<times> 'a \<Rightarrow> 'a \<times> 'a \<Rightarrow> 'a \<Rightarrow> 'a::euclidean_space"
+definition\<^marker>\<open>tag important\<close> interval_bij :: "'a \<times> 'a \<Rightarrow> 'a \<times> 'a \<Rightarrow> 'a \<Rightarrow> 'a::euclidean_space"
where "interval_bij =
(\<lambda>(a, b) (u, v) x. (\<Sum>i\<in>Basis. (u\<bullet>i + (x\<bullet>i - a\<bullet>i) / (b\<bullet>i - a\<bullet>i) * (v\<bullet>i - u\<bullet>i)) *\<^sub>R i))"
@@ -470,7 +470,7 @@
qed
-subsection%unimportant \<open>Some slightly ad hoc lemmas I use below\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Some slightly ad hoc lemmas I use below\<close>
lemma segment_vertical:
fixes a :: "real^2"
--- a/src/HOL/Analysis/Finite_Cartesian_Product.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Finite_Cartesian_Product.thy Fri Apr 12 22:09:25 2019 +0200
@@ -13,7 +13,7 @@
"HOL-Library.FuncSet"
begin
-subsection%unimportant \<open>Finite Cartesian products, with indexing and lambdas\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Finite Cartesian products, with indexing and lambdas\<close>
typedef ('a, 'b) vec = "UNIV :: ('b::finite \<Rightarrow> 'a) set"
morphisms vec_nth vec_lambda ..
@@ -165,7 +165,7 @@
by (simp add: image_comp o_def vec_nth_inverse)
qed
-subsection%unimportant \<open>Group operations and class instances\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Group operations and class instances\<close>
instantiation vec :: (zero, finite) zero
begin
@@ -230,7 +230,7 @@
by standard (simp_all add: vec_eq_iff)
-subsection%unimportant\<open>Basic componentwise operations on vectors\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Basic componentwise operations on vectors\<close>
instantiation vec :: (times, finite) times
begin
@@ -294,15 +294,15 @@
subsection \<open>Real vector space\<close>
-instantiation%unimportant vec :: (real_vector, finite) real_vector
+instantiation\<^marker>\<open>tag unimportant\<close> vec :: (real_vector, finite) real_vector
begin
-definition%important "scaleR \<equiv> (\<lambda> r x. (\<chi> i. scaleR r (x$i)))"
+definition\<^marker>\<open>tag important\<close> "scaleR \<equiv> (\<lambda> r x. (\<chi> i. scaleR r (x$i)))"
lemma vector_scaleR_component [simp]: "(scaleR r x)$i = scaleR r (x$i)"
unfolding scaleR_vec_def by simp
-instance%unimportant
+instance\<^marker>\<open>tag unimportant\<close>
by standard (simp_all add: vec_eq_iff scaleR_left_distrib scaleR_right_distrib)
end
@@ -310,15 +310,15 @@
subsection \<open>Topological space\<close>
-instantiation%unimportant vec :: (topological_space, finite) topological_space
+instantiation\<^marker>\<open>tag unimportant\<close> vec :: (topological_space, finite) topological_space
begin
-definition%important [code del]:
+definition\<^marker>\<open>tag important\<close> [code del]:
"open (S :: ('a ^ 'b) set) \<longleftrightarrow>
(\<forall>x\<in>S. \<exists>A. (\<forall>i. open (A i) \<and> x$i \<in> A i) \<and>
(\<forall>y. (\<forall>i. y$i \<in> A i) \<longrightarrow> y \<in> S))"
-instance%unimportant proof
+instance\<^marker>\<open>tag unimportant\<close> proof
show "open (UNIV :: ('a ^ 'b) set)"
unfolding open_vec_def by auto
next
@@ -418,7 +418,7 @@
then show "\<exists>T. open T \<and> a \<in> T \<and> T \<subseteq> (\<lambda>x. x $ i) ` S" by - (rule exI)
qed
-instance%unimportant vec :: (perfect_space, finite) perfect_space
+instance\<^marker>\<open>tag unimportant\<close> vec :: (perfect_space, finite) perfect_space
proof
fix x :: "'a ^ 'b" show "\<not> open {x}"
proof
@@ -433,29 +433,29 @@
subsection \<open>Metric space\<close>
(* TODO: Product of uniform spaces and compatibility with metric_spaces! *)
-instantiation%unimportant vec :: (metric_space, finite) dist
+instantiation\<^marker>\<open>tag unimportant\<close> vec :: (metric_space, finite) dist
begin
-definition%important
+definition\<^marker>\<open>tag important\<close>
"dist x y = L2_set (\<lambda>i. dist (x$i) (y$i)) UNIV"
instance ..
end
-instantiation%unimportant vec :: (metric_space, finite) uniformity_dist
+instantiation\<^marker>\<open>tag unimportant\<close> vec :: (metric_space, finite) uniformity_dist
begin
-definition%important [code del]:
+definition\<^marker>\<open>tag important\<close> [code del]:
"(uniformity :: (('a^'b::_) \<times> ('a^'b::_)) filter) =
(INF e\<in>{0 <..}. principal {(x, y). dist x y < e})"
-instance%unimportant
+instance\<^marker>\<open>tag unimportant\<close>
by standard (rule uniformity_vec_def)
end
declare uniformity_Abort[where 'a="'a :: metric_space ^ 'b :: finite", code]
-instantiation%unimportant vec :: (metric_space, finite) metric_space
+instantiation\<^marker>\<open>tag unimportant\<close> vec :: (metric_space, finite) metric_space
begin
proposition dist_vec_nth_le: "dist (x $ i) (y $ i) \<le> dist x y"
@@ -557,7 +557,7 @@
then show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < r" ..
qed
-instance%unimportant vec :: (complete_space, finite) complete_space
+instance\<^marker>\<open>tag unimportant\<close> vec :: (complete_space, finite) complete_space
proof
fix X :: "nat \<Rightarrow> 'a ^ 'b" assume "Cauchy X"
have "\<And>i. (\<lambda>n. X n $ i) \<longlonglongrightarrow> lim (\<lambda>n. X n $ i)"
@@ -572,14 +572,14 @@
subsection \<open>Normed vector space\<close>
-instantiation%unimportant vec :: (real_normed_vector, finite) real_normed_vector
+instantiation\<^marker>\<open>tag unimportant\<close> vec :: (real_normed_vector, finite) real_normed_vector
begin
-definition%important "norm x = L2_set (\<lambda>i. norm (x$i)) UNIV"
+definition\<^marker>\<open>tag important\<close> "norm x = L2_set (\<lambda>i. norm (x$i)) UNIV"
-definition%important "sgn (x::'a^'b) = scaleR (inverse (norm x)) x"
+definition\<^marker>\<open>tag important\<close> "sgn (x::'a^'b) = scaleR (inverse (norm x)) x"
-instance%unimportant proof
+instance\<^marker>\<open>tag unimportant\<close> proof
fix a :: real and x y :: "'a ^ 'b"
show "norm x = 0 \<longleftrightarrow> x = 0"
unfolding norm_vec_def
@@ -609,8 +609,8 @@
fixes x :: "'a::real_normed_vector^'n"
assumes "\<And>i. norm(x$i) \<le> norm(y$i)"
shows "norm x \<le> norm y"
- unfolding%unimportant norm_vec_def
- by%unimportant (rule L2_set_mono) (auto simp: assms)
+ unfolding norm_vec_def
+ by (rule L2_set_mono) (auto simp: assms)
lemma component_le_norm_cart: "\<bar>x$i\<bar> \<le> norm x"
apply (simp add: norm_vec_def)
@@ -638,12 +638,12 @@
subsection \<open>Inner product space\<close>
-instantiation%unimportant vec :: (real_inner, finite) real_inner
+instantiation\<^marker>\<open>tag unimportant\<close> vec :: (real_inner, finite) real_inner
begin
-definition%important "inner x y = sum (\<lambda>i. inner (x$i) (y$i)) UNIV"
+definition\<^marker>\<open>tag important\<close> "inner x y = sum (\<lambda>i. inner (x$i) (y$i)) UNIV"
-instance%unimportant proof
+instance\<^marker>\<open>tag unimportant\<close> proof
fix r :: real and x y z :: "'a ^ 'b"
show "inner x y = inner y x"
unfolding inner_vec_def
@@ -672,7 +672,7 @@
text \<open>Vectors pointing along a single axis.\<close>
-definition%important "axis k x = (\<chi> i. if i = k then x else 0)"
+definition\<^marker>\<open>tag important\<close> "axis k x = (\<chi> i. if i = k then x else 0)"
lemma axis_nth [simp]: "axis i x $ i = x"
unfolding axis_def by simp
@@ -696,12 +696,12 @@
lemma inner_axis': "inner(axis i y) x = inner y (x $ i)"
by (simp add: inner_axis inner_commute)
-instantiation%unimportant vec :: (euclidean_space, finite) euclidean_space
+instantiation\<^marker>\<open>tag unimportant\<close> vec :: (euclidean_space, finite) euclidean_space
begin
-definition%important "Basis = (\<Union>i. \<Union>u\<in>Basis. {axis i u})"
+definition\<^marker>\<open>tag important\<close> "Basis = (\<Union>i. \<Union>u\<in>Basis. {axis i u})"
-instance%unimportant proof
+instance\<^marker>\<open>tag unimportant\<close> proof
show "(Basis :: ('a ^ 'b) set) \<noteq> {}"
unfolding Basis_vec_def by simp
next
@@ -777,7 +777,7 @@
shows "(axis_index (axis u 1 :: real^'n)) = (u::'n)"
by (simp add: axis_eq_axis axis_index_def)
-subsection%unimportant \<open>A naive proof procedure to lift really trivial arithmetic stuff from the basis of the vector space\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>A naive proof procedure to lift really trivial arithmetic stuff from the basis of the vector space\<close>
lemma sum_cong_aux:
"(\<And>x. x \<in> A \<Longrightarrow> f x = g x) \<Longrightarrow> sum f A = sum g A"
@@ -851,13 +851,13 @@
lemma cond_component: "(if b then x else y)$i = (if b then x$i else y$i)" by vector
-lemmas%unimportant vector_component =
+lemmas\<^marker>\<open>tag unimportant\<close> vector_component =
vec_component vector_add_component vector_mult_component
vector_smult_component vector_minus_component vector_uminus_component
vector_scaleR_component cond_component
-subsection%unimportant \<open>Some frequently useful arithmetic lemmas over vectors\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Some frequently useful arithmetic lemmas over vectors\<close>
instance vec :: (semigroup_mult, finite) semigroup_mult
by standard (vector mult.assoc)
@@ -996,31 +996,31 @@
text\<open>Matrix notation. NB: an MxN matrix is of type \<^typ>\<open>'a^'n^'m\<close>, not \<^typ>\<open>'a^'m^'n\<close>\<close>
-definition%important map_matrix::"('a \<Rightarrow> 'b) \<Rightarrow> (('a, 'i::finite)vec, 'j::finite) vec \<Rightarrow> (('b, 'i)vec, 'j) vec" where
+definition\<^marker>\<open>tag important\<close> map_matrix::"('a \<Rightarrow> 'b) \<Rightarrow> (('a, 'i::finite)vec, 'j::finite) vec \<Rightarrow> (('b, 'i)vec, 'j) vec" where
"map_matrix f x = (\<chi> i j. f (x $ i $ j))"
lemma nth_map_matrix[simp]: "map_matrix f x $ i $ j = f (x $ i $ j)"
by (simp add: map_matrix_def)
-definition%important matrix_matrix_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'p^'n \<Rightarrow> 'a ^ 'p ^'m"
+definition\<^marker>\<open>tag important\<close> matrix_matrix_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'p^'n \<Rightarrow> 'a ^ 'p ^'m"
(infixl "**" 70)
where "m ** m' == (\<chi> i j. sum (\<lambda>k. ((m$i)$k) * ((m'$k)$j)) (UNIV :: 'n set)) ::'a ^ 'p ^'m"
-definition%important matrix_vector_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n \<Rightarrow> 'a ^ 'm"
+definition\<^marker>\<open>tag important\<close> matrix_vector_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n \<Rightarrow> 'a ^ 'm"
(infixl "*v" 70)
where "m *v x \<equiv> (\<chi> i. sum (\<lambda>j. ((m$i)$j) * (x$j)) (UNIV ::'n set)) :: 'a^'m"
-definition%important vector_matrix_mult :: "'a ^ 'm \<Rightarrow> ('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n "
+definition\<^marker>\<open>tag important\<close> vector_matrix_mult :: "'a ^ 'm \<Rightarrow> ('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n "
(infixl "v*" 70)
where "v v* m == (\<chi> j. sum (\<lambda>i. ((m$i)$j) * (v$i)) (UNIV :: 'm set)) :: 'a^'n"
-definition%unimportant "(mat::'a::zero => 'a ^'n^'n) k = (\<chi> i j. if i = j then k else 0)"
-definition%unimportant transpose where
+definition\<^marker>\<open>tag unimportant\<close> "(mat::'a::zero => 'a ^'n^'n) k = (\<chi> i j. if i = j then k else 0)"
+definition\<^marker>\<open>tag unimportant\<close> transpose where
"(transpose::'a^'n^'m \<Rightarrow> 'a^'m^'n) A = (\<chi> i j. ((A$j)$i))"
-definition%unimportant "(row::'m => 'a ^'n^'m \<Rightarrow> 'a ^'n) i A = (\<chi> j. ((A$i)$j))"
-definition%unimportant "(column::'n =>'a^'n^'m =>'a^'m) j A = (\<chi> i. ((A$i)$j))"
-definition%unimportant "rows(A::'a^'n^'m) = { row i A | i. i \<in> (UNIV :: 'm set)}"
-definition%unimportant "columns(A::'a^'n^'m) = { column i A | i. i \<in> (UNIV :: 'n set)}"
+definition\<^marker>\<open>tag unimportant\<close> "(row::'m => 'a ^'n^'m \<Rightarrow> 'a ^'n) i A = (\<chi> j. ((A$i)$j))"
+definition\<^marker>\<open>tag unimportant\<close> "(column::'n =>'a^'n^'m =>'a^'m) j A = (\<chi> i. ((A$i)$j))"
+definition\<^marker>\<open>tag unimportant\<close> "rows(A::'a^'n^'m) = { row i A | i. i \<in> (UNIV :: 'm set)}"
+definition\<^marker>\<open>tag unimportant\<close> "columns(A::'a^'n^'m) = { column i A | i. i \<in> (UNIV :: 'n set)}"
lemma times0_left [simp]: "(0::'a::semiring_1^'n^'m) ** (A::'a ^'p^'n) = 0"
by (simp add: matrix_matrix_mult_def zero_vec_def)
@@ -1151,7 +1151,7 @@
text\<open>Correspondence between matrices and linear operators.\<close>
-definition%important matrix :: "('a::{plus,times, one, zero}^'m \<Rightarrow> 'a ^ 'n) \<Rightarrow> 'a^'m^'n"
+definition\<^marker>\<open>tag important\<close> matrix :: "('a::{plus,times, one, zero}^'m \<Rightarrow> 'a ^ 'n) \<Rightarrow> 'a^'m^'n"
where "matrix f = (\<chi> i j. (f(axis j 1))$i)"
lemma matrix_id_mat_1: "matrix id = mat 1"
@@ -1204,10 +1204,10 @@
subsection\<open>Inverse matrices (not necessarily square)\<close>
-definition%important
+definition\<^marker>\<open>tag important\<close>
"invertible(A::'a::semiring_1^'n^'m) \<longleftrightarrow> (\<exists>A'::'a^'m^'n. A ** A' = mat 1 \<and> A' ** A = mat 1)"
-definition%important
+definition\<^marker>\<open>tag important\<close>
"matrix_inv(A:: 'a::semiring_1^'n^'m) =
(SOME A'::'a^'m^'n. A ** A' = mat 1 \<and> A' ** A = mat 1)"
--- a/src/HOL/Analysis/Finite_Product_Measure.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Finite_Product_Measure.thy Fri Apr 12 22:09:25 2019 +0200
@@ -15,7 +15,7 @@
lemma case_prod_const: "(\<lambda>(i, j). c) = (\<lambda>_. c)"
by auto
-subsection%unimportant \<open>More about Function restricted by \<^const>\<open>extensional\<close>\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>More about Function restricted by \<^const>\<open>extensional\<close>\<close>
definition
"merge I J = (\<lambda>(x, y) i. if i \<in> I then x i else if i \<in> J then y i else undefined)"
@@ -111,12 +111,12 @@
subsection \<open>Finite product spaces\<close>
-definition%important prod_emb where
+definition\<^marker>\<open>tag important\<close> prod_emb where
"prod_emb I M K X = (\<lambda>x. restrict x K) -` X \<inter> (\<Pi>\<^sub>E i\<in>I. space (M i))"
lemma prod_emb_iff:
"f \<in> prod_emb I M K X \<longleftrightarrow> f \<in> extensional I \<and> (restrict f K \<in> X) \<and> (\<forall>i\<in>I. f i \<in> space (M i))"
- unfolding%unimportant prod_emb_def PiE_def by auto
+ unfolding prod_emb_def PiE_def by auto
lemma
shows prod_emb_empty[simp]: "prod_emb M L K {} = {}"
@@ -153,13 +153,13 @@
"F \<subseteq> G \<Longrightarrow> prod_emb A M B F \<subseteq> prod_emb A M B G"
by (auto simp: prod_emb_def)
-definition%important PiM :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i \<Rightarrow> 'a) measure" where
+definition\<^marker>\<open>tag important\<close> PiM :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i \<Rightarrow> 'a) measure" where
"PiM I M = extend_measure (\<Pi>\<^sub>E i\<in>I. space (M i))
{(J, X). (J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}
(\<lambda>(J, X). prod_emb I M J (\<Pi>\<^sub>E j\<in>J. X j))
(\<lambda>(J, X). \<Prod>j\<in>J \<union> {i\<in>I. emeasure (M i) (space (M i)) \<noteq> 1}. if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))"
-definition%important prod_algebra :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i \<Rightarrow> 'a) set set" where
+definition\<^marker>\<open>tag important\<close> prod_algebra :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i \<Rightarrow> 'a) set set" where
"prod_algebra I M = (\<lambda>(J, X). prod_emb I M J (\<Pi>\<^sub>E j\<in>J. X j)) `
{(J, X). (J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}"
@@ -825,14 +825,14 @@
by (auto intro!: eq)
qed (auto simp: sets_PiM)
-locale%unimportant product_sigma_finite =
+locale\<^marker>\<open>tag unimportant\<close> product_sigma_finite =
fixes M :: "'i \<Rightarrow> 'a measure"
assumes sigma_finite_measures: "\<And>i. sigma_finite_measure (M i)"
-sublocale%unimportant product_sigma_finite \<subseteq> M?: sigma_finite_measure "M i" for i
+sublocale\<^marker>\<open>tag unimportant\<close> product_sigma_finite \<subseteq> M?: sigma_finite_measure "M i" for i
by (rule sigma_finite_measures)
-locale%unimportant finite_product_sigma_finite = product_sigma_finite M for M :: "'i \<Rightarrow> 'a measure" +
+locale\<^marker>\<open>tag unimportant\<close> finite_product_sigma_finite = product_sigma_finite M for M :: "'i \<Rightarrow> 'a measure" +
fixes I :: "'i set"
assumes finite_index: "finite I"
@@ -991,7 +991,7 @@
"0 \<le> f (\<lambda>k. undefined) \<Longrightarrow> integral\<^sup>N (Pi\<^sub>M {} M) f = f (\<lambda>k. undefined)"
by (simp add: PiM_empty nn_integral_count_space_finite max.absorb2)
-lemma%important (in product_sigma_finite) distr_merge:
+lemma\<^marker>\<open>tag important\<close> (in product_sigma_finite) distr_merge:
assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
shows "distr (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \<union> J) M) (merge I J) = Pi\<^sub>M (I \<union> J) M"
(is "?D = ?P")
--- a/src/HOL/Analysis/Function_Topology.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Function_Topology.thy Fri Apr 12 22:09:25 2019 +0200
@@ -74,7 +74,7 @@
So, we use the first formulation, which moreover seems to give rise to more straightforward proofs.
\<close>
-definition%important product_topology::"('i \<Rightarrow> ('a topology)) \<Rightarrow> ('i set) \<Rightarrow> (('i \<Rightarrow> 'a) topology)"
+definition\<^marker>\<open>tag important\<close> product_topology::"('i \<Rightarrow> ('a topology)) \<Rightarrow> ('i set) \<Rightarrow> (('i \<Rightarrow> 'a) topology)"
where "product_topology T I =
topology_generated_by {(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. openin (T i) (X i)) \<and> finite {i. X i \<noteq> topspace (T i)}}"
@@ -548,12 +548,12 @@
qed
-subsubsection%important \<open>Powers of a single topological space as a topological space, using type classes\<close>
+subsubsection\<^marker>\<open>tag important\<close> \<open>Powers of a single topological space as a topological space, using type classes\<close>
instantiation "fun" :: (type, topological_space) topological_space
begin
-definition%important open_fun_def:
+definition\<^marker>\<open>tag important\<close> open_fun_def:
"open U = openin (product_topology (\<lambda>i. euclidean) UNIV) U"
instance proof
@@ -653,7 +653,7 @@
by (intro continuous_intros) auto
qed
-subsubsection%important \<open>Topological countability for product spaces\<close>
+subsubsection\<^marker>\<open>tag important\<close> \<open>Topological countability for product spaces\<close>
text \<open>The next two lemmas are useful to prove first or second countability
of product spaces, but they have more to do with countability and could
@@ -897,10 +897,10 @@
instantiation "fun" :: (countable, metric_space) metric_space
begin
-definition%important dist_fun_def:
+definition\<^marker>\<open>tag important\<close> dist_fun_def:
"dist x y = (\<Sum>n. (1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1)"
-definition%important uniformity_fun_def:
+definition\<^marker>\<open>tag important\<close> uniformity_fun_def:
"(uniformity::(('a \<Rightarrow> 'b) \<times> ('a \<Rightarrow> 'b)) filter) = (INF e\<in>{0<..}. principal {(x, y). dist (x::('a\<Rightarrow>'b)) y < e})"
text \<open>Except for the first one, the auxiliary lemmas below are only useful when proving the
--- a/src/HOL/Analysis/Further_Topology.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Further_Topology.thy Fri Apr 12 22:09:25 2019 +0200
@@ -2249,7 +2249,7 @@
assumes contf: "continuous_on S f" and "open S"
and injf: "inj_on f S" and "S \<noteq> {}"
shows "DIM('a) \<le> DIM('b)"
- using%unimportant invariance_of_dimension_subspaces [of UNIV S UNIV f] assms
+ using invariance_of_dimension_subspaces [of UNIV S UNIV f] assms
by auto
@@ -2260,7 +2260,7 @@
and injf: "inj_on f S"
shows "dim S \<le> dim T"
apply (rule invariance_of_dimension_subspaces [of S S _ f])
- using%unimportant assms by (auto simp: subspace_affine)
+ using assms by (auto simp: subspace_affine)
lemma invariance_of_dimension_convex_domain:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
@@ -2364,8 +2364,8 @@
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "open S" "continuous_on S f" "DIM('b) \<le> DIM('a)" "inj_on f S"
shows "S homeomorphic (f ` S)"
- using%unimportant invariance_of_domain_homeomorphism [OF assms]
- by%unimportant (meson homeomorphic_def)
+ using invariance_of_domain_homeomorphism [OF assms]
+ by (meson homeomorphic_def)
lemma continuous_image_subset_interior:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
@@ -3375,7 +3375,7 @@
corollary covering_space_square_punctured_plane:
"covering_space (- {0}) (\<lambda>z::complex. z^2) (- {0})"
- by%unimportant (simp add: covering_space_power_punctured_plane)
+ by (simp add: covering_space_power_punctured_plane)
proposition covering_space_exp_punctured_plane:
@@ -4200,7 +4200,7 @@
text\<open>This doesn't have a standard name. Kuratowski uses ``contractible with respect to \<open>[S\<^sup>1]\<close>''
while Whyburn uses ``property b''. It's closely related to unicoherence.\<close>
-definition%important Borsukian where
+definition\<^marker>\<open>tag important\<close> Borsukian where
"Borsukian S \<equiv>
\<forall>f. continuous_on S f \<and> f ` S \<subseteq> (- {0::complex})
\<longrightarrow> (\<exists>a. homotopic_with_canon (\<lambda>h. True) S (- {0}) f (\<lambda>x. a))"
@@ -4815,7 +4815,7 @@
subsection\<open>Unicoherence (closed)\<close>
-definition%important unicoherent where
+definition\<^marker>\<open>tag important\<close> unicoherent where
"unicoherent U \<equiv>
\<forall>S T. connected S \<and> connected T \<and> S \<union> T = U \<and>
closedin (top_of_set U) S \<and> closedin (top_of_set U) T
@@ -5006,16 +5006,16 @@
corollary contractible_imp_unicoherent:
fixes U :: "'a::euclidean_space set"
assumes "contractible U" shows "unicoherent U"
- by%unimportant (simp add: Borsukian_imp_unicoherent assms contractible_imp_Borsukian)
+ by (simp add: Borsukian_imp_unicoherent assms contractible_imp_Borsukian)
corollary convex_imp_unicoherent:
fixes U :: "'a::euclidean_space set"
assumes "convex U" shows "unicoherent U"
- by%unimportant (simp add: Borsukian_imp_unicoherent assms convex_imp_Borsukian)
+ by (simp add: Borsukian_imp_unicoherent assms convex_imp_Borsukian)
text\<open>If the type class constraint can be relaxed, I don't know how!\<close>
corollary unicoherent_UNIV: "unicoherent (UNIV :: 'a :: euclidean_space set)"
- by%unimportant (simp add: convex_imp_unicoherent)
+ by (simp add: convex_imp_unicoherent)
lemma unicoherent_monotone_image_compact:
--- a/src/HOL/Analysis/Gamma_Function.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Gamma_Function.thy Fri Apr 12 22:09:25 2019 +0200
@@ -245,7 +245,7 @@
immediately from the definition.
\<close>
-definition%important Gamma_series :: "('a :: {banach,real_normed_field}) \<Rightarrow> nat \<Rightarrow> 'a" where
+definition\<^marker>\<open>tag important\<close> Gamma_series :: "('a :: {banach,real_normed_field}) \<Rightarrow> nat \<Rightarrow> 'a" where
"Gamma_series z n = fact n * exp (z * of_real (ln (of_nat n))) / pochhammer z (n+1)"
definition Gamma_series' :: "('a :: {banach,real_normed_field}) \<Rightarrow> nat \<Rightarrow> 'a" where
@@ -317,10 +317,10 @@
function to the inverse of the Gamma function, and from that to the Gamma function itself.
\<close>
-definition%important ln_Gamma_series :: "('a :: {banach,real_normed_field,ln}) \<Rightarrow> nat \<Rightarrow> 'a" where
+definition\<^marker>\<open>tag important\<close> ln_Gamma_series :: "('a :: {banach,real_normed_field,ln}) \<Rightarrow> nat \<Rightarrow> 'a" where
"ln_Gamma_series z n = z * ln (of_nat n) - ln z - (\<Sum>k=1..n. ln (z / of_nat k + 1))"
-definition%unimportant ln_Gamma_series' :: "('a :: {banach,real_normed_field,ln}) \<Rightarrow> nat \<Rightarrow> 'a" where
+definition\<^marker>\<open>tag unimportant\<close> ln_Gamma_series' :: "('a :: {banach,real_normed_field,ln}) \<Rightarrow> nat \<Rightarrow> 'a" where
"ln_Gamma_series' z n =
- euler_mascheroni*z - ln z + (\<Sum>k=1..n. z / of_nat n - ln (z / of_nat k + 1))"
@@ -614,12 +614,12 @@
by (rule uniformly_convergent_imp_convergent,
rule uniformly_summable_deriv_ln_Gamma[of z "norm z/2"]) simp_all
-definition%important Polygamma :: "nat \<Rightarrow> ('a :: {real_normed_field,banach}) \<Rightarrow> 'a" where
+definition\<^marker>\<open>tag important\<close> Polygamma :: "nat \<Rightarrow> ('a :: {real_normed_field,banach}) \<Rightarrow> 'a" where
"Polygamma n z = (if n = 0 then
(\<Sum>k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) - euler_mascheroni else
(-1)^Suc n * fact n * (\<Sum>k. inverse ((z + of_nat k)^Suc n)))"
-abbreviation%important Digamma :: "('a :: {real_normed_field,banach}) \<Rightarrow> 'a" where
+abbreviation\<^marker>\<open>tag important\<close> Digamma :: "('a :: {real_normed_field,banach}) \<Rightarrow> 'a" where
"Digamma \<equiv> Polygamma 0"
lemma Digamma_def:
@@ -1021,7 +1021,7 @@
the reals and for the complex numbers with a minimal amount of proof duplication.
\<close>
-class%unimportant Gamma = real_normed_field + complete_space +
+class\<^marker>\<open>tag unimportant\<close> Gamma = real_normed_field + complete_space +
fixes rGamma :: "'a \<Rightarrow> 'a"
assumes rGamma_eq_zero_iff_aux: "rGamma z = 0 \<longleftrightarrow> (\<exists>n. z = - of_nat n)"
assumes differentiable_rGamma_aux1:
@@ -1259,12 +1259,12 @@
"isCont f z \<Longrightarrow> f z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> isCont (\<lambda>x. Gamma (f x)) z"
by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_Gamma]])
-subsection%unimportant \<open>The complex Gamma function\<close>
-
-instantiation%unimportant complex :: Gamma
+subsection\<^marker>\<open>tag unimportant\<close> \<open>The complex Gamma function\<close>
+
+instantiation\<^marker>\<open>tag unimportant\<close> complex :: Gamma
begin
-definition%unimportant rGamma_complex :: "complex \<Rightarrow> complex" where
+definition\<^marker>\<open>tag unimportant\<close> rGamma_complex :: "complex \<Rightarrow> complex" where
"rGamma_complex z = lim (rGamma_series z)"
lemma rGamma_series_complex_converges:
@@ -1299,7 +1299,7 @@
thus "?thesis1" "?thesis2" by blast+
qed
-context%unimportant
+context\<^marker>\<open>tag unimportant\<close>
begin
(* TODO: duplication *)
@@ -1526,7 +1526,7 @@
-subsection%unimportant \<open>The real Gamma function\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>The real Gamma function\<close>
lemma rGamma_series_real:
"eventually (\<lambda>n. rGamma_series x n = Re (rGamma_series (of_real x) n)) sequentially"
@@ -1544,7 +1544,7 @@
finally show "rGamma_series x n = Re (rGamma_series (of_real x) n)" ..
qed
-instantiation%unimportant real :: Gamma
+instantiation\<^marker>\<open>tag unimportant\<close> real :: Gamma
begin
definition "rGamma_real x = Re (rGamma (of_real x :: complex))"
@@ -1807,7 +1807,7 @@
integers, where the Gamma function is not defined).
\<close>
-context%unimportant
+context\<^marker>\<open>tag unimportant\<close>
fixes G :: "real \<Rightarrow> real"
assumes G_1: "G 1 = 1"
assumes G_plus1: "x > 0 \<Longrightarrow> G (x + 1) = x * G x"
@@ -2389,7 +2389,7 @@
end
-subsection%unimportant \<open>Limits and residues\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Limits and residues\<close>
text \<open>
The inverse of the Gamma function has simple zeros:
@@ -2537,7 +2537,7 @@
"Gamma_series_Weierstrass z n =
exp (-euler_mascheroni * z) / z * (\<Prod>k=1..n. exp (z / of_nat k) / (1 + z / of_nat k))"
-definition%unimportant
+definition\<^marker>\<open>tag unimportant\<close>
rGamma_series_Weierstrass :: "'a :: {banach,real_normed_field} \<Rightarrow> nat \<Rightarrow> 'a" where
"rGamma_series_Weierstrass z n =
exp (euler_mascheroni * z) * z * (\<Prod>k=1..n. (1 + z / of_nat k) * exp (-z / of_nat k))"
--- a/src/HOL/Analysis/Great_Picard.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Great_Picard.thy Fri Apr 12 22:09:25 2019 +0200
@@ -788,7 +788,7 @@
-subsubsection%important\<open>Montel's theorem\<close>
+subsubsection\<^marker>\<open>tag important\<close>\<open>Montel's theorem\<close>
text\<open>a sequence of holomorphic functions uniformly bounded
on compact subsets of an open set S has a subsequence that converges to a
@@ -1780,8 +1780,8 @@
assumes M: "open M" "z \<in> M" and holf: "f holomorphic_on (M - {z})"
and non: "\<And>l. \<not> (f \<longlongrightarrow> l) (at z)" "\<And>l. \<not> ((inverse \<circ> f) \<longlongrightarrow> l) (at z)"
obtains a where "- {a} \<subseteq> f ` (M - {z})"
- apply%unimportant (simp add: subset_iff image_iff)
- by%unimportant (metis great_Picard [OF M _ holf] non)
+ apply (simp add: subset_iff image_iff)
+ by (metis great_Picard [OF M _ holf] non)
corollary great_Picard_infinite:
--- a/src/HOL/Analysis/Harmonic_Numbers.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Harmonic_Numbers.thy Fri Apr 12 22:09:25 2019 +0200
@@ -19,7 +19,7 @@
subsection \<open>The Harmonic numbers\<close>
-definition%important harm :: "nat \<Rightarrow> 'a :: real_normed_field" where
+definition\<^marker>\<open>tag important\<close> harm :: "nat \<Rightarrow> 'a :: real_normed_field" where
"harm n = (\<Sum>k=1..n. inverse (of_nat k))"
lemma harm_altdef: "harm n = (\<Sum>k<n. inverse (of_nat (Suc k)))"
@@ -156,7 +156,7 @@
thus ?thesis by (subst (asm) convergent_Suc_iff)
qed
-lemma%important euler_mascheroni_LIMSEQ:
+lemma\<^marker>\<open>tag important\<close> euler_mascheroni_LIMSEQ:
"(\<lambda>n. harm n - ln (of_nat n) :: real) \<longlonglongrightarrow> euler_mascheroni"
unfolding euler_mascheroni_def
by (simp add: convergent_LIMSEQ_iff [symmetric] euler_mascheroni_convergent)
@@ -250,7 +250,7 @@
qed
-subsection%unimportant \<open>Bounds on the Euler-Mascheroni constant\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Bounds on the Euler-Mascheroni constant\<close>
(* TODO: Move? *)
lemma ln_inverse_approx_le:
--- a/src/HOL/Analysis/Homeomorphism.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Homeomorphism.thy Fri Apr 12 22:09:25 2019 +0200
@@ -852,7 +852,7 @@
assumes "0 < r" and b: "b \<in> sphere a r"
and "affine T" and affS: "aff_dim T + 1 = DIM('a)"
shows "(sphere a r - {b}) homeomorphic T"
- using%unimportant homeomorphic_punctured_affine_sphere_affine [of r b a UNIV T] assms by%unimportant auto
+ using homeomorphic_punctured_affine_sphere_affine [of r b a UNIV T] assms by auto
corollary homeomorphic_punctured_sphere_hyperplane:
fixes a :: "'a :: euclidean_space"
@@ -1270,7 +1270,7 @@
subsection\<open>Covering spaces and lifting results for them\<close>
-definition%important covering_space
+definition\<^marker>\<open>tag important\<close> covering_space
:: "'a::topological_space set \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b::topological_space set \<Rightarrow> bool"
where
"covering_space c p S \<equiv>
@@ -2130,8 +2130,8 @@
and "path h2" and pih2: "path_image h2 \<subseteq> C" and ph2: "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h2 t) = g2 t"
and h1h2: "pathstart h1 = pathstart h2"
shows "pathfinish h1 = pathfinish h2"
- using%unimportant covering_space_lift_homotopic_paths [OF assms] homotopic_paths_imp_pathfinish
- by%unimportant blast
+ using covering_space_lift_homotopic_paths [OF assms] homotopic_paths_imp_pathfinish
+ by blast
corollary covering_space_lift_homotopic_path:
--- a/src/HOL/Analysis/Homotopy.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Homotopy.thy Fri Apr 12 22:09:25 2019 +0200
@@ -8,7 +8,7 @@
imports Path_Connected Continuum_Not_Denumerable Product_Topology
begin
-definition%important homotopic_with
+definition\<^marker>\<open>tag important\<close> homotopic_with
where
"homotopic_with P X Y f g \<equiv>
(\<exists>h. continuous_map (prod_topology (subtopology euclideanreal {0..1}) X) Y h \<and>
@@ -50,10 +50,10 @@
apply (simp add: t)
done
-subsection%unimportant\<open>Trivial properties\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Trivial properties\<close>
text \<open>We often want to just localize the ending function equality or whatever.\<close>
-text%important \<open>%whitespace\<close>
+text\<^marker>\<open>tag important\<close> \<open>%whitespace\<close>
proposition homotopic_with:
assumes "\<And>h k. (\<And>x. x \<in> topspace X \<Longrightarrow> h x = k x) \<Longrightarrow> (P h \<longleftrightarrow> P k)"
shows "homotopic_with P X Y p q \<longleftrightarrow>
@@ -516,7 +516,7 @@
subsection\<open>Homotopy of paths, maintaining the same endpoints\<close>
-definition%important homotopic_paths :: "['a set, real \<Rightarrow> 'a, real \<Rightarrow> 'a::topological_space] \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> homotopic_paths :: "['a set, real \<Rightarrow> 'a, real \<Rightarrow> 'a::topological_space] \<Rightarrow> bool"
where
"homotopic_paths s p q \<equiv>
homotopic_with_canon (\<lambda>r. pathstart r = pathstart p \<and> pathfinish r = pathfinish p) {0..1} s p q"
@@ -673,7 +673,7 @@
subsection\<open>Group properties for homotopy of paths\<close>
-text%important\<open>So taking equivalence classes under homotopy would give the fundamental group\<close>
+text\<^marker>\<open>tag important\<close>\<open>So taking equivalence classes under homotopy would give the fundamental group\<close>
proposition homotopic_paths_rid:
"\<lbrakk>path p; path_image p \<subseteq> s\<rbrakk> \<Longrightarrow> homotopic_paths s (p +++ linepath (pathfinish p) (pathfinish p)) p"
@@ -736,7 +736,7 @@
subsection\<open>Homotopy of loops without requiring preservation of endpoints\<close>
-definition%important homotopic_loops :: "'a::topological_space set \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> bool" where
+definition\<^marker>\<open>tag important\<close> homotopic_loops :: "'a::topological_space set \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> bool" where
"homotopic_loops s p q \<equiv>
homotopic_with_canon (\<lambda>r. pathfinish r = pathstart r) {0..1} s p q"
@@ -971,7 +971,7 @@
qed
-subsection%unimportant\<open>Homotopy of "nearby" function, paths and loops\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Homotopy of "nearby" function, paths and loops\<close>
lemma homotopic_with_linear:
fixes f g :: "_ \<Rightarrow> 'b::real_normed_vector"
@@ -1173,9 +1173,9 @@
subsection\<open>Simply connected sets\<close>
-text%important\<open>defined as "all loops are homotopic (as loops)\<close>
-
-definition%important simply_connected where
+text\<^marker>\<open>tag important\<close>\<open>defined as "all loops are homotopic (as loops)\<close>
+
+definition\<^marker>\<open>tag important\<close> simply_connected where
"simply_connected S \<equiv>
\<forall>p q. path p \<and> pathfinish p = pathstart p \<and> path_image p \<subseteq> S \<and>
path q \<and> pathfinish q = pathstart q \<and> path_image q \<subseteq> S
@@ -1357,7 +1357,7 @@
subsection\<open>Contractible sets\<close>
-definition%important contractible where
+definition\<^marker>\<open>tag important\<close> contractible where
"contractible S \<equiv> \<exists>a. homotopic_with_canon (\<lambda>x. True) S S id (\<lambda>x. a)"
proposition contractible_imp_simply_connected:
@@ -1598,7 +1598,7 @@
subsection\<open>Local versions of topological properties in general\<close>
-definition%important locally :: "('a::topological_space set \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> locally :: "('a::topological_space set \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool"
where
"locally P S \<equiv>
\<forall>w x. openin (top_of_set S) w \<and> x \<in> w
@@ -2975,7 +2975,7 @@
by metis
qed
-subsection%unimportant\<open>Components, continuity, openin, closedin\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Components, continuity, openin, closedin\<close>
lemma continuous_on_components_gen:
fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
@@ -3308,7 +3308,7 @@
subsection\<open>Retracts, in a general sense, preserve (co)homotopic triviality)\<close>
-locale%important Retracts =
+locale\<^marker>\<open>tag important\<close> Retracts =
fixes s h t k
assumes conth: "continuous_on s h"
and imh: "h ` s = t"
@@ -3474,7 +3474,7 @@
subsection\<open>Homotopy equivalence of topological spaces.\<close>
-definition%important homotopy_equivalent_space
+definition\<^marker>\<open>tag important\<close> homotopy_equivalent_space
(infix "homotopy'_equivalent'_space" 50)
where "X homotopy_equivalent_space Y \<equiv>
(\<exists>f g. continuous_map X Y f \<and>
@@ -3897,7 +3897,7 @@
qed
-abbreviation%important homotopy_eqv :: "'a::topological_space set \<Rightarrow> 'b::topological_space set \<Rightarrow> bool"
+abbreviation\<^marker>\<open>tag important\<close> homotopy_eqv :: "'a::topological_space set \<Rightarrow> 'b::topological_space set \<Rightarrow> bool"
(infix "homotopy'_eqv" 50)
where "S homotopy_eqv T \<equiv> top_of_set S homotopy_equivalent_space top_of_set T"
@@ -4144,7 +4144,7 @@
by (metis homeomorphic_contractible_eq)
-subsection%unimportant\<open>Misc other results\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Misc other results\<close>
lemma bounded_connected_Compl_real:
fixes S :: "real set"
@@ -4194,7 +4194,7 @@
qed
-subsection%unimportant\<open>Some Uncountable Sets\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Some Uncountable Sets\<close>
lemma uncountable_closed_segment:
fixes a :: "'a::real_normed_vector"
@@ -4333,7 +4333,7 @@
by (simp add: arc_imp_simple_path assms simple_path_image_uncountable)
-subsection%unimportant\<open> Some simple positive connection theorems\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open> Some simple positive connection theorems\<close>
proposition path_connected_convex_diff_countable:
fixes U :: "'a::euclidean_space set"
@@ -4598,9 +4598,9 @@
-subsection%unimportant \<open>Self-homeomorphisms shuffling points about\<close>
-
-subsubsection%unimportant\<open>The theorem \<open>homeomorphism_moving_points_exists\<close>\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Self-homeomorphisms shuffling points about\<close>
+
+subsubsection\<^marker>\<open>tag unimportant\<close>\<open>The theorem \<open>homeomorphism_moving_points_exists\<close>\<close>
lemma homeomorphism_moving_point_1:
fixes a :: "'a::euclidean_space"
@@ -4730,7 +4730,7 @@
done
qed
-corollary%unimportant homeomorphism_moving_point_2:
+corollary\<^marker>\<open>tag unimportant\<close> homeomorphism_moving_point_2:
fixes a :: "'a::euclidean_space"
assumes "affine T" "a \<in> T" and u: "u \<in> ball a r \<inter> T" and v: "v \<in> ball a r \<inter> T"
obtains f g where "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f g"
@@ -4758,7 +4758,7 @@
qed
-corollary%unimportant homeomorphism_moving_point_3:
+corollary\<^marker>\<open>tag unimportant\<close> homeomorphism_moving_point_3:
fixes a :: "'a::euclidean_space"
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"
@@ -4836,7 +4836,7 @@
qed
-proposition%unimportant homeomorphism_moving_point:
+proposition\<^marker>\<open>tag unimportant\<close> homeomorphism_moving_point:
fixes a :: "'a::euclidean_space"
assumes ope: "openin (top_of_set (affine hull S)) S"
and "S \<subseteq> T"
@@ -4992,7 +4992,7 @@
qed
qed
-proposition%unimportant homeomorphism_moving_points_exists:
+proposition\<^marker>\<open>tag unimportant\<close> homeomorphism_moving_points_exists:
fixes S :: "'a::euclidean_space set"
assumes 2: "2 \<le> DIM('a)" "open S" "connected S" "S \<subseteq> T" "finite K"
and KS: "\<And>i. i \<in> K \<Longrightarrow> x i \<in> S \<and> y i \<in> S"
@@ -5021,7 +5021,7 @@
using homeomorphism_moving_points_exists_gen [OF \<open>finite K\<close> KS pw _ ope S] that by fastforce
qed
-subsubsection%unimportant\<open>The theorem \<open>homeomorphism_grouping_points_exists\<close>\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close>\<open>The theorem \<open>homeomorphism_grouping_points_exists\<close>\<close>
lemma homeomorphism_grouping_point_1:
fixes a::real and c::real
@@ -5258,7 +5258,7 @@
qed
qed
-proposition%unimportant homeomorphism_grouping_points_exists:
+proposition\<^marker>\<open>tag unimportant\<close> 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. (\<not> (f x = x \<and> g x = x))} \<subseteq> S"
@@ -5337,7 +5337,7 @@
qed
-proposition%unimportant homeomorphism_grouping_points_exists_gen:
+proposition\<^marker>\<open>tag unimportant\<close> homeomorphism_grouping_points_exists_gen:
fixes S :: "'a::euclidean_space set"
assumes opeU: "openin (top_of_set S) U"
and opeS: "openin (top_of_set (affine hull S)) S"
--- a/src/HOL/Analysis/Improper_Integral.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Improper_Integral.thy Fri Apr 12 22:09:25 2019 +0200
@@ -9,7 +9,7 @@
text\<open>The definition here only really makes sense for an elementary set.
We just use compact intervals in applications below.\<close>
-definition%important equiintegrable_on (infixr "equiintegrable'_on" 46)
+definition\<^marker>\<open>tag important\<close> equiintegrable_on (infixr "equiintegrable'_on" 46)
where "F equiintegrable_on I \<equiv>
(\<forall>f \<in> F. f integrable_on I) \<and>
(\<forall>e > 0. \<exists>\<gamma>. gauge \<gamma> \<and>
--- a/src/HOL/Analysis/Infinite_Products.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Infinite_Products.thy Fri Apr 12 22:09:25 2019 +0200
@@ -9,7 +9,7 @@
imports Topology_Euclidean_Space Complex_Transcendental
begin
-subsection%unimportant \<open>Preliminaries\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Preliminaries\<close>
lemma sum_le_prod:
fixes f :: "'a \<Rightarrow> 'b :: linordered_semidom"
@@ -54,19 +54,19 @@
subsection\<open>Definitions and basic properties\<close>
-definition%important raw_has_prod :: "[nat \<Rightarrow> 'a::{t2_space, comm_semiring_1}, nat, 'a] \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> raw_has_prod :: "[nat \<Rightarrow> 'a::{t2_space, comm_semiring_1}, nat, 'a] \<Rightarrow> bool"
where "raw_has_prod f M p \<equiv> (\<lambda>n. \<Prod>i\<le>n. f (i+M)) \<longlonglongrightarrow> p \<and> p \<noteq> 0"
text\<open>The nonzero and zero cases, as in \emph{Complex Analysis} by Joseph Bak and Donald J.Newman, page 241\<close>
-text%important \<open>%whitespace\<close>
-definition%important
+text\<^marker>\<open>tag important\<close> \<open>%whitespace\<close>
+definition\<^marker>\<open>tag important\<close>
has_prod :: "(nat \<Rightarrow> 'a::{t2_space, comm_semiring_1}) \<Rightarrow> 'a \<Rightarrow> bool" (infixr "has'_prod" 80)
where "f has_prod p \<equiv> raw_has_prod f 0 p \<or> (\<exists>i q. p = 0 \<and> f i = 0 \<and> raw_has_prod f (Suc i) q)"
-definition%important convergent_prod :: "(nat \<Rightarrow> 'a :: {t2_space,comm_semiring_1}) \<Rightarrow> bool" where
+definition\<^marker>\<open>tag important\<close> convergent_prod :: "(nat \<Rightarrow> 'a :: {t2_space,comm_semiring_1}) \<Rightarrow> bool" where
"convergent_prod f \<equiv> \<exists>M p. raw_has_prod f M p"
-definition%important prodinf :: "(nat \<Rightarrow> 'a::{t2_space, comm_semiring_1}) \<Rightarrow> 'a"
+definition\<^marker>\<open>tag important\<close> prodinf :: "(nat \<Rightarrow> 'a::{t2_space, comm_semiring_1}) \<Rightarrow> 'a"
(binder "\<Prod>" 10)
where "prodinf f = (THE p. f has_prod p)"
@@ -148,7 +148,7 @@
subsection\<open>Absolutely convergent products\<close>
-definition%important abs_convergent_prod :: "(nat \<Rightarrow> _) \<Rightarrow> bool" where
+definition\<^marker>\<open>tag important\<close> abs_convergent_prod :: "(nat \<Rightarrow> _) \<Rightarrow> bool" where
"abs_convergent_prod f \<longleftrightarrow> convergent_prod (\<lambda>i. 1 + norm (f i - 1))"
lemma abs_convergent_prodI:
@@ -222,7 +222,7 @@
by (rule_tac x=0 in exI) auto
qed
-lemma%important convergent_prod_iff_convergent:
+lemma\<^marker>\<open>tag important\<close> convergent_prod_iff_convergent:
fixes f :: "nat \<Rightarrow> 'a :: {topological_semigroup_mult,t2_space,idom}"
assumes "\<And>i. f i \<noteq> 0"
shows "convergent_prod f \<longleftrightarrow> convergent (\<lambda>n. \<Prod>i\<le>n. f i) \<and> lim (\<lambda>n. \<Prod>i\<le>n. f i) \<noteq> 0"
@@ -398,7 +398,7 @@
thus ?thesis by eventually_elim auto
qed
-subsection%unimportant \<open>Ignoring initial segments\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Ignoring initial segments\<close>
lemma convergent_prod_offset:
assumes "convergent_prod (\<lambda>n. f (n + m))"
@@ -450,7 +450,7 @@
using raw_has_prod_def that by blast
qed
-corollary%unimportant convergent_prod_ignore_initial_segment:
+corollary\<^marker>\<open>tag unimportant\<close> convergent_prod_ignore_initial_segment:
fixes f :: "nat \<Rightarrow> 'a :: real_normed_field"
assumes "convergent_prod f"
shows "convergent_prod (\<lambda>n. f (n + m))"
@@ -461,14 +461,14 @@
apply (auto simp add: raw_has_prod_def add_ac)
done
-corollary%unimportant convergent_prod_ignore_nonzero_segment:
+corollary\<^marker>\<open>tag unimportant\<close> convergent_prod_ignore_nonzero_segment:
fixes f :: "nat \<Rightarrow> 'a :: real_normed_field"
assumes f: "convergent_prod f" and nz: "\<And>i. i \<ge> M \<Longrightarrow> f i \<noteq> 0"
shows "\<exists>p. raw_has_prod f M p"
using convergent_prod_ignore_initial_segment [OF f]
by (metis convergent_LIMSEQ_iff convergent_prod_iff_convergent le_add_same_cancel2 nz prod_defs(1) zero_order(1))
-corollary%unimportant abs_convergent_prod_ignore_initial_segment:
+corollary\<^marker>\<open>tag unimportant\<close> abs_convergent_prod_ignore_initial_segment:
assumes "abs_convergent_prod f"
shows "abs_convergent_prod (\<lambda>n. f (n + m))"
using assms unfolding abs_convergent_prod_def
@@ -761,7 +761,7 @@
qed
qed
-corollary%unimportant has_prod_0:
+corollary\<^marker>\<open>tag unimportant\<close> has_prod_0:
fixes f :: "nat \<Rightarrow> 'a::{semidom,t2_space}"
assumes "\<And>n. f n = 1"
shows "f has_prod 1"
@@ -874,7 +874,7 @@
end
-subsection%unimportant \<open>Infinite products on ordered topological monoids\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Infinite products on ordered topological monoids\<close>
lemma LIMSEQ_prod_0:
fixes f :: "nat \<Rightarrow> 'a::{semidom,topological_space}"
@@ -1075,7 +1075,7 @@
using convergent_prod_def raw_has_prodI_bounded [OF assms] by blast
-subsection%unimportant \<open>Infinite products on topological spaces\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Infinite products on topological spaces\<close>
context
fixes f g :: "nat \<Rightarrow> 'a::{t2_space,topological_semigroup_mult,idom}"
@@ -1176,7 +1176,7 @@
end
-subsection%unimportant \<open>Infinite summability on real normed fields\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Infinite summability on real normed fields\<close>
context
fixes f :: "nat \<Rightarrow> 'a::real_normed_field"
@@ -1332,7 +1332,7 @@
by (simp add: ac_simps)
qed
-corollary%unimportant has_prod_iff_shift':
+corollary\<^marker>\<open>tag unimportant\<close> has_prod_iff_shift':
assumes "\<And>i. i < n \<Longrightarrow> f i \<noteq> 0"
shows "(\<lambda>i. f (i + n)) has_prod (a / (\<Prod>i<n. f i)) \<longleftrightarrow> f has_prod a"
by (simp add: assms has_prod_iff_shift)
@@ -1773,7 +1773,7 @@
using convergent_prod_def assms convergent_prod_iff_summable_complex by blast
-subsection%unimportant \<open>Embeddings from the reals into some complete real normed field\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Embeddings from the reals into some complete real normed field\<close>
lemma tendsto_eq_of_real_lim:
assumes "(\<lambda>n. of_real (f n) :: 'a::{complete_space,real_normed_field}) \<longlonglongrightarrow> q"
--- a/src/HOL/Analysis/Infinite_Set_Sum.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Infinite_Set_Sum.thy Fri Apr 12 22:09:25 2019 +0200
@@ -74,14 +74,14 @@
-definition%important abs_summable_on ::
+definition\<^marker>\<open>tag important\<close> abs_summable_on ::
"('a \<Rightarrow> 'b :: {banach, second_countable_topology}) \<Rightarrow> 'a set \<Rightarrow> bool"
(infix "abs'_summable'_on" 50)
where
"f abs_summable_on A \<longleftrightarrow> integrable (count_space A) f"
-definition%important infsetsum ::
+definition\<^marker>\<open>tag important\<close> infsetsum ::
"('a \<Rightarrow> 'b :: {banach, second_countable_topology}) \<Rightarrow> 'a set \<Rightarrow> 'b"
where
"infsetsum f A = lebesgue_integral (count_space A) f"
--- a/src/HOL/Analysis/Inner_Product.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Inner_Product.thy Fri Apr 12 22:09:25 2019 +0200
@@ -390,7 +390,7 @@
subsection \<open>Gradient derivative\<close>
-definition%important
+definition\<^marker>\<open>tag important\<close>
gderiv ::
"['a::real_inner \<Rightarrow> real, 'a, 'a] \<Rightarrow> bool"
("(GDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60)
--- a/src/HOL/Analysis/Integral_Test.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Integral_Test.thy Fri Apr 12 22:09:25 2019 +0200
@@ -19,7 +19,7 @@
\<close>
(* TODO: continuous_in \<rightarrow> integrable_on *)
-locale%important antimono_fun_sum_integral_diff =
+locale\<^marker>\<open>tag important\<close> antimono_fun_sum_integral_diff =
fixes f :: "real \<Rightarrow> real"
assumes dec: "\<And>x y. x \<ge> 0 \<Longrightarrow> x \<le> y \<Longrightarrow> f x \<ge> f y"
assumes nonneg: "\<And>x. x \<ge> 0 \<Longrightarrow> f x \<ge> 0"
--- a/src/HOL/Analysis/Interval_Integral.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Interval_Integral.thy Fri Apr 12 22:09:25 2019 +0200
@@ -129,7 +129,7 @@
(* TODO: in this definition, it would be more natural if einterval a b included a and b when
they are real. *)
-definition%important interval_lebesgue_integral :: "real measure \<Rightarrow> ereal \<Rightarrow> ereal \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> 'a::{banach, second_countable_topology}" where
+definition\<^marker>\<open>tag important\<close> interval_lebesgue_integral :: "real measure \<Rightarrow> ereal \<Rightarrow> ereal \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> 'a::{banach, second_countable_topology}" where
"interval_lebesgue_integral M a b f =
(if a \<le> b then (LINT x:einterval a b|M. f x) else - (LINT x:einterval b a|M. f x))"
@@ -140,7 +140,7 @@
translations
"LINT x=a..b|M. f" == "CONST interval_lebesgue_integral M a b (\<lambda>x. f)"
-definition%important interval_lebesgue_integrable :: "real measure \<Rightarrow> ereal \<Rightarrow> ereal \<Rightarrow> (real \<Rightarrow> 'a::{banach, second_countable_topology}) \<Rightarrow> bool" where
+definition\<^marker>\<open>tag important\<close> interval_lebesgue_integrable :: "real measure \<Rightarrow> ereal \<Rightarrow> ereal \<Rightarrow> (real \<Rightarrow> 'a::{banach, second_countable_topology}) \<Rightarrow> bool" where
"interval_lebesgue_integrable M a b f =
(if a \<le> b then set_integrable M (einterval a b) f else set_integrable M (einterval b a) f)"
@@ -301,7 +301,7 @@
proposition interval_integrable_to_infinity_eq: "(interval_lebesgue_integrable M a \<infinity> f) =
(set_integrable M {a<..} f)"
- unfolding%unimportant interval_lebesgue_integrable_def by auto
+ unfolding interval_lebesgue_integrable_def by auto
subsection\<open>Basic properties of integration over an interval wrt lebesgue measure\<close>
@@ -1083,8 +1083,8 @@
fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
shows "interval_lebesgue_integrable lborel a b f \<Longrightarrow> a \<le> b \<Longrightarrow>
norm (LBINT t=a..b. f t) \<le> LBINT t=a..b. norm (f t)"
- using%unimportant integral_norm_bound[of lborel "\<lambda>x. indicator (einterval a b) x *\<^sub>R f x"]
- by%unimportant (auto simp: interval_lebesgue_integral_def interval_lebesgue_integrable_def set_lebesgue_integral_def)
+ using integral_norm_bound[of lborel "\<lambda>x. indicator (einterval a b) x *\<^sub>R f x"]
+ by (auto simp: interval_lebesgue_integral_def interval_lebesgue_integrable_def set_lebesgue_integral_def)
proposition interval_integral_norm2:
"interval_lebesgue_integrable lborel a b f \<Longrightarrow>
--- a/src/HOL/Analysis/Jordan_Curve.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Jordan_Curve.thy Fri Apr 12 22:09:25 2019 +0200
@@ -219,7 +219,7 @@
qed
-theorem%unimportant Jordan_curve:
+theorem\<^marker>\<open>tag unimportant\<close> Jordan_curve:
fixes c :: "real \<Rightarrow> complex"
assumes "simple_path c" and loop: "pathfinish c = pathstart c"
obtains inner outer where
@@ -389,7 +389,7 @@
qed
-corollary%unimportant Jordan_disconnected:
+corollary\<^marker>\<open>tag unimportant\<close> Jordan_disconnected:
fixes c :: "real \<Rightarrow> complex"
assumes "simple_path c" "pathfinish c = pathstart c"
shows "\<not> connected(- path_image c)"
--- a/src/HOL/Analysis/L2_Norm.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/L2_Norm.thy Fri Apr 12 22:09:25 2019 +0200
@@ -10,7 +10,7 @@
section \<open>L2 Norm\<close>
-definition%important L2_set :: "('a \<Rightarrow> real) \<Rightarrow> 'a set \<Rightarrow> real" where
+definition\<^marker>\<open>tag important\<close> L2_set :: "('a \<Rightarrow> real) \<Rightarrow> 'a set \<Rightarrow> real" where
"L2_set f A = sqrt (\<Sum>i\<in>A. (f i)\<^sup>2)"
lemma L2_set_cong:
--- a/src/HOL/Analysis/Lebesgue_Measure.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Lebesgue_Measure.thy Fri Apr 12 22:09:25 2019 +0200
@@ -40,16 +40,16 @@
Every right-continuous and nondecreasing function gives rise to a measure on the reals:
\<close>
-definition%important interval_measure :: "(real \<Rightarrow> real) \<Rightarrow> real measure" where
+definition\<^marker>\<open>tag important\<close> interval_measure :: "(real \<Rightarrow> real) \<Rightarrow> real measure" where
"interval_measure F =
extend_measure UNIV {(a, b). a \<le> b} (\<lambda>(a, b). {a<..b}) (\<lambda>(a, b). ennreal (F b - F a))"
-lemma%important emeasure_interval_measure_Ioc:
+lemma\<^marker>\<open>tag important\<close> emeasure_interval_measure_Ioc:
assumes "a \<le> b"
assumes mono_F: "\<And>x y. x \<le> y \<Longrightarrow> F x \<le> F y"
assumes right_cont_F : "\<And>a. continuous (at_right a) F"
shows "emeasure (interval_measure F) {a<..b} = F b - F a"
-proof%unimportant (rule extend_measure_caratheodory_pair[OF interval_measure_def \<open>a \<le> b\<close>])
+proof (rule extend_measure_caratheodory_pair[OF interval_measure_def \<open>a \<le> b\<close>])
show "semiring_of_sets UNIV {{a<..b} |a b :: real. a \<le> b}"
proof (unfold_locales, safe)
fix a b c d :: real assume *: "a \<le> b" "c \<le> d"
@@ -319,7 +319,7 @@
emeasure (interval_measure F) {a <.. b} = (if a \<le> b then F b - F a else 0)"
using emeasure_interval_measure_Ioc[of a b F] by auto
-lemma%important sets_interval_measure [simp, measurable_cong]:
+lemma\<^marker>\<open>tag important\<close> sets_interval_measure [simp, measurable_cong]:
"sets (interval_measure F) = sets borel"
apply (simp add: sets_extend_measure interval_measure_def borel_sigma_sets_Ioc)
apply (rule sigma_sets_eqI)
@@ -371,7 +371,7 @@
(auto simp: continuous_on_ennreal continuous_on_diff cont_F continuous_on_const)
qed (rule trivial_limit_at_left_real)
-lemma%important sigma_finite_interval_measure:
+lemma\<^marker>\<open>tag important\<close> sigma_finite_interval_measure:
assumes mono_F: "\<And>x y. x \<le> y \<Longrightarrow> F x \<le> F y"
assumes right_cont_F : "\<And>a. continuous (at_right a) F"
shows "sigma_finite_measure (interval_measure F)"
@@ -382,13 +382,13 @@
subsection \<open>Lebesgue-Borel measure\<close>
-definition%important lborel :: "('a :: euclidean_space) measure" where
+definition\<^marker>\<open>tag important\<close> lborel :: "('a :: euclidean_space) measure" where
"lborel = distr (\<Pi>\<^sub>M b\<in>Basis. interval_measure (\<lambda>x. x)) borel (\<lambda>f. \<Sum>b\<in>Basis. f b *\<^sub>R b)"
-abbreviation%important lebesgue :: "'a::euclidean_space measure"
+abbreviation\<^marker>\<open>tag important\<close> lebesgue :: "'a::euclidean_space measure"
where "lebesgue \<equiv> completion lborel"
-abbreviation%important lebesgue_on :: "'a set \<Rightarrow> 'a::euclidean_space measure"
+abbreviation\<^marker>\<open>tag important\<close> lebesgue_on :: "'a set \<Rightarrow> 'a::euclidean_space measure"
where "lebesgue_on \<Omega> \<equiv> restrict_space (completion lborel) \<Omega>"
lemma
@@ -409,7 +409,7 @@
shows "f \<in> measurable (lebesgue_on S) M \<longleftrightarrow> g \<in> measurable (lebesgue_on S) M"
by (metis (mono_tags, lifting) IntD1 assms measurable_cong_simp space_restrict_space)
-text%unimportant \<open>Measurability of continuous functions\<close>
+text\<^marker>\<open>tag unimportant\<close> \<open>Measurability of continuous functions\<close>
lemma continuous_imp_measurable_on_sets_lebesgue:
assumes f: "continuous_on S f" and S: "S \<in> sets lebesgue"
@@ -460,10 +460,10 @@
lemma emeasure_lborel_Icc_eq: "emeasure lborel {l .. u} = ennreal (if l \<le> u then u - l else 0)"
by simp
-lemma%important emeasure_lborel_cbox[simp]:
+lemma\<^marker>\<open>tag important\<close> emeasure_lborel_cbox[simp]:
assumes [simp]: "\<And>b. b \<in> Basis \<Longrightarrow> l \<bullet> b \<le> u \<bullet> b"
shows "emeasure lborel (cbox l u) = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
-proof%unimportant -
+proof -
have "(\<lambda>x. \<Prod>b\<in>Basis. indicator {l\<bullet>b .. u\<bullet>b} (x \<bullet> b) :: ennreal) = indicator (cbox l u)"
by (auto simp: fun_eq_iff cbox_def split: split_indicator)
then have "emeasure lborel (cbox l u) = (\<integral>\<^sup>+x. (\<Prod>b\<in>Basis. indicator {l\<bullet>b .. u\<bullet>b} (x \<bullet> b)) \<partial>lborel)"
@@ -654,12 +654,12 @@
subsection \<open>Affine transformation on the Lebesgue-Borel\<close>
-lemma%important lborel_eqI:
+lemma\<^marker>\<open>tag important\<close> lborel_eqI:
fixes M :: "'a::euclidean_space measure"
assumes emeasure_eq: "\<And>l u. (\<And>b. b \<in> Basis \<Longrightarrow> l \<bullet> b \<le> u \<bullet> b) \<Longrightarrow> emeasure M (box l u) = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
assumes sets_eq: "sets M = sets borel"
shows "lborel = M"
-proof%unimportant (rule measure_eqI_generator_eq)
+proof (rule measure_eqI_generator_eq)
let ?E = "range (\<lambda>(a, b). box a b::'a set)"
show "Int_stable ?E"
by (auto simp: Int_stable_def box_Int_box)
@@ -679,12 +679,12 @@
done }
qed
-lemma%important lborel_affine_euclidean:
+lemma\<^marker>\<open>tag important\<close> lborel_affine_euclidean:
fixes c :: "'a::euclidean_space \<Rightarrow> real" and t
defines "T x \<equiv> t + (\<Sum>j\<in>Basis. (c j * (x \<bullet> j)) *\<^sub>R j)"
assumes c: "\<And>j. j \<in> Basis \<Longrightarrow> c j \<noteq> 0"
shows "lborel = density (distr lborel borel T) (\<lambda>_. (\<Prod>j\<in>Basis. \<bar>c j\<bar>))" (is "_ = ?D")
-proof%unimportant (rule lborel_eqI)
+proof (rule lborel_eqI)
let ?B = "Basis :: 'a set"
fix l u assume le: "\<And>b. b \<in> ?B \<Longrightarrow> l \<bullet> b \<le> u \<bullet> b"
have [measurable]: "T \<in> borel \<rightarrow>\<^sub>M borel"
@@ -736,10 +736,10 @@
lborel_integrable_real_affine[of "\<lambda>x. f (t + c * x)" "1/c" "-t/c"]
by (auto simp add: field_simps)
-lemma%important lborel_integral_real_affine:
+lemma\<^marker>\<open>tag important\<close> lborel_integral_real_affine:
fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}" and c :: real
assumes c: "c \<noteq> 0" shows "(\<integral>x. f x \<partial> lborel) = \<bar>c\<bar> *\<^sub>R (\<integral>x. f (t + c * x) \<partial>lborel)"
-proof%unimportant cases
+proof cases
assume f[measurable]: "integrable lborel f" then show ?thesis
using c f f[THEN borel_measurable_integrable] f[THEN lborel_integrable_real_affine, of c t]
by (subst lborel_real_affine[OF c, of t])
@@ -901,9 +901,9 @@
interpretation lborel_pair: pair_sigma_finite lborel lborel ..
-lemma%important lborel_prod:
+lemma\<^marker>\<open>tag important\<close> lborel_prod:
"lborel \<Otimes>\<^sub>M lborel = (lborel :: ('a::euclidean_space \<times> 'b::euclidean_space) measure)"
-proof%unimportant (rule lborel_eqI[symmetric], clarify)
+proof (rule lborel_eqI[symmetric], clarify)
fix la ua :: 'a and lb ub :: 'b
assume lu: "\<And>a b. (a, b) \<in> Basis \<Longrightarrow> (la, lb) \<bullet> (a, b) \<le> (ua, ub) \<bullet> (a, b)"
have [simp]:
@@ -978,14 +978,14 @@
subsection \<open>Lebesgue measurable sets\<close>
-abbreviation%important lmeasurable :: "'a::euclidean_space set set"
+abbreviation\<^marker>\<open>tag important\<close> lmeasurable :: "'a::euclidean_space set set"
where
"lmeasurable \<equiv> fmeasurable lebesgue"
lemma not_measurable_UNIV [simp]: "UNIV \<notin> lmeasurable"
by (simp add: fmeasurable_def)
-lemma%important lmeasurable_iff_integrable:
+lemma\<^marker>\<open>tag important\<close> lmeasurable_iff_integrable:
"S \<in> lmeasurable \<longleftrightarrow> integrable lebesgue (indicator S :: 'a::euclidean_space \<Rightarrow> real)"
by (auto simp: fmeasurable_def integrable_iff_bounded borel_measurable_indicator_iff ennreal_indicator)
@@ -1039,7 +1039,7 @@
by (metis assms bounded_Un emeasure_bounded_finite emeasure_completion fmeasurableI main_part_null_part_Un)
-subsection%unimportant\<open>Translation preserves Lebesgue measure\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Translation preserves Lebesgue measure\<close>
lemma sigma_sets_image:
assumes S: "S \<in> sigma_sets \<Omega> M" and "M \<subseteq> Pow \<Omega>" "f ` \<Omega> = \<Omega>" "inj_on f \<Omega>"
@@ -1128,12 +1128,12 @@
lemma summable_iff_suminf_neq_top: "(\<And>n. f n \<ge> 0) \<Longrightarrow> \<not> summable f \<Longrightarrow> (\<Sum>i. ennreal (f i)) = top"
by (metis summable_suminf_not_top)
-proposition%important starlike_negligible_bounded_gmeasurable:
+proposition\<^marker>\<open>tag important\<close> starlike_negligible_bounded_gmeasurable:
fixes S :: "'a :: euclidean_space set"
assumes S: "S \<in> sets lebesgue" and "bounded S"
and eq1: "\<And>c x. \<lbrakk>(c *\<^sub>R x) \<in> S; 0 \<le> c; x \<in> S\<rbrakk> \<Longrightarrow> c = 1"
shows "S \<in> null_sets lebesgue"
-proof%unimportant -
+proof -
obtain M where "0 < M" "S \<subseteq> ball 0 M"
using \<open>bounded S\<close> by (auto dest: bounded_subset_ballD)
@@ -1252,10 +1252,10 @@
qed
qed
-lemma%important outer_regular_lborel:
+lemma\<^marker>\<open>tag important\<close> outer_regular_lborel:
assumes B: "B \<in> sets borel" and "0 < (e::real)"
obtains U where "open U" "B \<subseteq> U" "emeasure lborel (U - B) < e"
-proof%unimportant -
+proof -
obtain U where U: "open U" "B \<subseteq> U" and "emeasure lborel (U-B) \<le> e/2"
using outer_regular_lborel_le [OF B, of "e/2"] \<open>e > 0\<close>
by force
--- a/src/HOL/Analysis/Linear_Algebra.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Linear_Algebra.thy Fri Apr 12 22:09:25 2019 +0200
@@ -104,7 +104,7 @@
then show ?thesis by auto
qed
-subsection%unimportant \<open>More interesting properties of the norm\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>More interesting properties of the norm\<close>
unbundle inner_syntax
@@ -230,7 +230,7 @@
subsection \<open>Orthogonality\<close>
-definition%important (in real_inner) "orthogonal x y \<longleftrightarrow> x \<bullet> y = 0"
+definition\<^marker>\<open>tag important\<close> (in real_inner) "orthogonal x y \<longleftrightarrow> x \<bullet> y = 0"
context real_inner
begin
@@ -299,9 +299,9 @@
subsection \<open>Orthogonality of a transformation\<close>
-definition%important "orthogonal_transformation f \<longleftrightarrow> linear f \<and> (\<forall>v w. f v \<bullet> f w = v \<bullet> w)"
+definition\<^marker>\<open>tag important\<close> "orthogonal_transformation f \<longleftrightarrow> linear f \<and> (\<forall>v w. f v \<bullet> f w = v \<bullet> w)"
-lemma%unimportant orthogonal_transformation:
+lemma\<^marker>\<open>tag unimportant\<close> orthogonal_transformation:
"orthogonal_transformation f \<longleftrightarrow> linear f \<and> (\<forall>v. norm (f v) = norm v)"
unfolding orthogonal_transformation_def
apply auto
@@ -310,56 +310,56 @@
apply (simp add: dot_norm linear_add[symmetric])
done
-lemma%unimportant orthogonal_transformation_id [simp]: "orthogonal_transformation (\<lambda>x. x)"
+lemma\<^marker>\<open>tag unimportant\<close> orthogonal_transformation_id [simp]: "orthogonal_transformation (\<lambda>x. x)"
by (simp add: linear_iff orthogonal_transformation_def)
-lemma%unimportant orthogonal_orthogonal_transformation:
+lemma\<^marker>\<open>tag unimportant\<close> orthogonal_orthogonal_transformation:
"orthogonal_transformation f \<Longrightarrow> orthogonal (f x) (f y) \<longleftrightarrow> orthogonal x y"
by (simp add: orthogonal_def orthogonal_transformation_def)
-lemma%unimportant orthogonal_transformation_compose:
+lemma\<^marker>\<open>tag unimportant\<close> orthogonal_transformation_compose:
"\<lbrakk>orthogonal_transformation f; orthogonal_transformation g\<rbrakk> \<Longrightarrow> orthogonal_transformation(f \<circ> g)"
by (auto simp: orthogonal_transformation_def linear_compose)
-lemma%unimportant orthogonal_transformation_neg:
+lemma\<^marker>\<open>tag unimportant\<close> orthogonal_transformation_neg:
"orthogonal_transformation(\<lambda>x. -(f x)) \<longleftrightarrow> orthogonal_transformation f"
by (auto simp: orthogonal_transformation_def dest: linear_compose_neg)
-lemma%unimportant orthogonal_transformation_scaleR: "orthogonal_transformation f \<Longrightarrow> f (c *\<^sub>R v) = c *\<^sub>R f v"
+lemma\<^marker>\<open>tag unimportant\<close> orthogonal_transformation_scaleR: "orthogonal_transformation f \<Longrightarrow> f (c *\<^sub>R v) = c *\<^sub>R f v"
by (simp add: linear_iff orthogonal_transformation_def)
-lemma%unimportant orthogonal_transformation_linear:
+lemma\<^marker>\<open>tag unimportant\<close> orthogonal_transformation_linear:
"orthogonal_transformation f \<Longrightarrow> linear f"
by (simp add: orthogonal_transformation_def)
-lemma%unimportant orthogonal_transformation_inj:
+lemma\<^marker>\<open>tag unimportant\<close> orthogonal_transformation_inj:
"orthogonal_transformation f \<Longrightarrow> inj f"
unfolding orthogonal_transformation_def inj_on_def
by (metis vector_eq)
-lemma%unimportant orthogonal_transformation_surj:
+lemma\<^marker>\<open>tag unimportant\<close> orthogonal_transformation_surj:
"orthogonal_transformation f \<Longrightarrow> surj f"
for f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space"
by (simp add: linear_injective_imp_surjective orthogonal_transformation_inj orthogonal_transformation_linear)
-lemma%unimportant orthogonal_transformation_bij:
+lemma\<^marker>\<open>tag unimportant\<close> orthogonal_transformation_bij:
"orthogonal_transformation f \<Longrightarrow> bij f"
for f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space"
by (simp add: bij_def orthogonal_transformation_inj orthogonal_transformation_surj)
-lemma%unimportant orthogonal_transformation_inv:
+lemma\<^marker>\<open>tag unimportant\<close> orthogonal_transformation_inv:
"orthogonal_transformation f \<Longrightarrow> orthogonal_transformation (inv f)"
for f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space"
by (metis (no_types, hide_lams) bijection.inv_right bijection_def inj_linear_imp_inv_linear orthogonal_transformation orthogonal_transformation_bij orthogonal_transformation_inj)
-lemma%unimportant orthogonal_transformation_norm:
+lemma\<^marker>\<open>tag unimportant\<close> orthogonal_transformation_norm:
"orthogonal_transformation f \<Longrightarrow> norm (f x) = norm x"
by (metis orthogonal_transformation)
subsection \<open>Bilinear functions\<close>
-definition%important
+definition\<^marker>\<open>tag important\<close>
bilinear :: "('a::real_vector \<Rightarrow> 'b::real_vector \<Rightarrow> 'c::real_vector) \<Rightarrow> bool" where
"bilinear f \<longleftrightarrow> (\<forall>x. linear (\<lambda>y. f x y)) \<and> (\<forall>y. linear (\<lambda>x. f x y))"
@@ -417,7 +417,7 @@
subsection \<open>Adjoints\<close>
-definition%important adjoint :: "(('a::real_inner) \<Rightarrow> ('b::real_inner)) \<Rightarrow> 'b \<Rightarrow> 'a" where
+definition\<^marker>\<open>tag important\<close> adjoint :: "(('a::real_inner) \<Rightarrow> ('b::real_inner)) \<Rightarrow> 'b \<Rightarrow> 'a" where
"adjoint f = (SOME f'. \<forall>x y. f x \<bullet> y = x \<bullet> f' y)"
lemma adjoint_unique:
@@ -579,7 +579,7 @@
done
-subsection%unimportant \<open>Euclidean Spaces as Typeclass\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Euclidean Spaces as Typeclass\<close>
lemma independent_Basis: "independent Basis"
by (rule independent_Basis)
@@ -591,7 +591,7 @@
unfolding span_Basis ..
-subsection%unimportant \<open>Linearity and Bilinearity continued\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Linearity and Bilinearity continued\<close>
lemma linear_bounded:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
@@ -773,7 +773,7 @@
by (metis linear_imp_has_derivative differentiable_def)
-subsection%unimportant \<open>We continue\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>We continue\<close>
lemma independent_bound:
fixes S :: "'a::euclidean_space set"
@@ -1037,7 +1037,7 @@
subsection \<open>Infinity norm\<close>
-definition%important "infnorm (x::'a::euclidean_space) = Sup {\<bar>x \<bullet> b\<bar> |b. b \<in> Basis}"
+definition\<^marker>\<open>tag important\<close> "infnorm (x::'a::euclidean_space) = Sup {\<bar>x \<bullet> b\<bar> |b. b \<in> Basis}"
lemma infnorm_set_image:
fixes x :: "'a::euclidean_space"
@@ -1227,7 +1227,7 @@
subsection \<open>Collinearity\<close>
-definition%important collinear :: "'a::real_vector set \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> collinear :: "'a::real_vector set \<Rightarrow> bool"
where "collinear S \<longleftrightarrow> (\<exists>u. \<forall>x \<in> S. \<forall> y \<in> S. \<exists>c. x - y = c *\<^sub>R u)"
lemma collinear_alt:
@@ -1565,7 +1565,7 @@
by (rule_tac U=U in that) (auto simp: span_Un)
qed
-corollary%unimportant orthogonal_extension_strong:
+corollary\<^marker>\<open>tag unimportant\<close> orthogonal_extension_strong:
fixes S :: "'a::euclidean_space set"
assumes S: "pairwise orthogonal S"
obtains U where "U \<inter> (insert 0 S) = {}" "pairwise orthogonal (S \<union> U)"
@@ -1648,7 +1648,7 @@
qed
-proposition%unimportant orthogonal_to_subspace_exists_gen:
+proposition\<^marker>\<open>tag unimportant\<close> orthogonal_to_subspace_exists_gen:
fixes S :: "'a :: euclidean_space set"
assumes "span S \<subset> span T"
obtains x where "x \<noteq> 0" "x \<in> span T" "\<And>y. y \<in> span S \<Longrightarrow> orthogonal x y"
@@ -1694,7 +1694,7 @@
qed
qed
-corollary%unimportant orthogonal_to_subspace_exists:
+corollary\<^marker>\<open>tag unimportant\<close> orthogonal_to_subspace_exists:
fixes S :: "'a :: euclidean_space set"
assumes "dim S < DIM('a)"
obtains x where "x \<noteq> 0" "\<And>y. y \<in> span S \<Longrightarrow> orthogonal x y"
@@ -1706,7 +1706,7 @@
by (auto simp: span_UNIV)
qed
-corollary%unimportant orthogonal_to_vector_exists:
+corollary\<^marker>\<open>tag unimportant\<close> orthogonal_to_vector_exists:
fixes x :: "'a :: euclidean_space"
assumes "2 \<le> DIM('a)"
obtains y where "y \<noteq> 0" "orthogonal x y"
@@ -1717,7 +1717,7 @@
by (rule orthogonal_to_subspace_exists) (simp add: orthogonal_commute span_base that)
qed
-proposition%unimportant orthogonal_subspace_decomp_exists:
+proposition\<^marker>\<open>tag unimportant\<close> orthogonal_subspace_decomp_exists:
fixes S :: "'a :: euclidean_space set"
obtains y z
where "y \<in> span S"
--- a/src/HOL/Analysis/Lipschitz.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Lipschitz.thy Fri Apr 12 22:09:25 2019 +0200
@@ -8,11 +8,11 @@
imports Borel_Space
begin
-definition%important lipschitz_on
+definition\<^marker>\<open>tag important\<close> lipschitz_on
where "lipschitz_on C U f \<longleftrightarrow> (0 \<le> C \<and> (\<forall>x \<in> U. \<forall>y\<in>U. dist (f x) (f y) \<le> C * dist x y))"
bundle lipschitz_syntax begin
-notation%important lipschitz_on ("_-lipschitz'_on" [1000])
+notation\<^marker>\<open>tag important\<close> lipschitz_on ("_-lipschitz'_on" [1000])
end
bundle no_lipschitz_syntax begin
no_notation lipschitz_on ("_-lipschitz'_on" [1000])
@@ -155,7 +155,7 @@
qed fact
-subsubsection%unimportant \<open>Structural introduction rules\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Structural introduction rules\<close>
named_theorems lipschitz_intros "structural introduction rules for Lipschitz controls"
@@ -481,7 +481,7 @@
subsection \<open>Local Lipschitz continuity (uniform for a family of functions)\<close>
-definition%important local_lipschitz::
+definition\<^marker>\<open>tag important\<close> local_lipschitz::
"'a::metric_space set \<Rightarrow> 'b::metric_space set \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c::metric_space) \<Rightarrow> bool"
where
"local_lipschitz T X f \<equiv> \<forall>x \<in> X. \<forall>t \<in> T.
--- a/src/HOL/Analysis/Measurable.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Measurable.thy Fri Apr 12 22:09:25 2019 +0200
@@ -64,10 +64,10 @@
attribute_setup measurable_cong = Measurable.cong_thm_attr
"add congurence rules to measurability prover"
-method_setup%important measurable = \<open> Scan.lift (Scan.succeed (METHOD o Measurable.measurable_tac)) \<close>
+method_setup\<^marker>\<open>tag important\<close> measurable = \<open> Scan.lift (Scan.succeed (METHOD o Measurable.measurable_tac)) \<close>
"measurability prover"
-simproc_setup%important measurable ("A \<in> sets M" | "f \<in> measurable M N") = \<open>K Measurable.simproc\<close>
+simproc_setup\<^marker>\<open>tag important\<close> measurable ("A \<in> sets M" | "f \<in> measurable M N") = \<open>K Measurable.simproc\<close>
setup \<open>
Global_Theory.add_thms_dynamic (\<^binding>\<open>measurable\<close>, Measurable.get_all)
@@ -380,7 +380,7 @@
unfolding pred_def
by (auto intro!: sets.sets_Collect_countable_All' sets.sets_Collect_countable_Ex' assms)
-subsection%unimportant \<open>Measurability for (co)inductive predicates\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Measurability for (co)inductive predicates\<close>
lemma measurable_bot[measurable]: "bot \<in> measurable M (count_space UNIV)"
by (simp add: bot_fun_def)
--- a/src/HOL/Analysis/Measure_Space.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Measure_Space.thy Fri Apr 12 22:09:25 2019 +0200
@@ -11,7 +11,7 @@
Measurable "HOL-Library.Extended_Nonnegative_Real"
begin
-subsection%unimportant "Relate extended reals and the indicator function"
+subsection\<^marker>\<open>tag unimportant\<close> "Relate extended reals and the indicator function"
lemma suminf_cmult_indicator:
fixes f :: "nat \<Rightarrow> ennreal"
@@ -59,7 +59,7 @@
is also used to represent sigma algebras (with an arbitrary emeasure).
\<close>
-subsection%unimportant "Extend binary sets"
+subsection\<^marker>\<open>tag unimportant\<close> "Extend binary sets"
lemma LIMSEQ_binaryset:
assumes f: "f {} = 0"
@@ -91,7 +91,7 @@
shows "f {} = 0 \<Longrightarrow> (\<Sum>n. f (binaryset A B n)) = f A + f B"
by (metis binaryset_sums sums_unique)
-subsection%unimportant \<open>Properties of a premeasure \<^term>\<open>\<mu>\<close>\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Properties of a premeasure \<^term>\<open>\<mu>\<close>\<close>
text \<open>
The definitions for \<^const>\<open>positive\<close> and \<^const>\<open>countably_additive\<close> should be here, by they are
@@ -442,7 +442,7 @@
using empty_continuous_imp_continuous_from_below[OF f fin] cont
by blast
-subsection%unimportant \<open>Properties of \<^const>\<open>emeasure\<close>\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Properties of \<^const>\<open>emeasure\<close>\<close>
lemma emeasure_positive: "positive (sets M) (emeasure M)"
by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def)
@@ -881,7 +881,7 @@
subsection \<open>\<open>\<mu>\<close>-null sets\<close>
-definition%important null_sets :: "'a measure \<Rightarrow> 'a set set" where
+definition\<^marker>\<open>tag important\<close> null_sets :: "'a measure \<Rightarrow> 'a set set" where
"null_sets M = {N\<in>sets M. emeasure M N = 0}"
lemma null_setsD1[dest]: "A \<in> null_sets M \<Longrightarrow> emeasure M A = 0"
@@ -989,7 +989,7 @@
subsection \<open>The almost everywhere filter (i.e.\ quantifier)\<close>
-definition%important ae_filter :: "'a measure \<Rightarrow> 'a filter" where
+definition\<^marker>\<open>tag important\<close> ae_filter :: "'a measure \<Rightarrow> 'a filter" where
"ae_filter M = (INF N\<in>null_sets M. principal (space M - N))"
abbreviation almost_everywhere :: "'a measure \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" where
@@ -1265,7 +1265,7 @@
subsection \<open>\<open>\<sigma>\<close>-finite Measures\<close>
-locale%important sigma_finite_measure =
+locale\<^marker>\<open>tag important\<close> sigma_finite_measure =
fixes M :: "'a measure"
assumes sigma_finite_countable:
"\<exists>A::'a set set. countable A \<and> A \<subseteq> sets M \<and> (\<Union>A) = space M \<and> (\<forall>a\<in>A. emeasure M a \<noteq> \<infinity>)"
@@ -1387,7 +1387,7 @@
subsection \<open>Measure space induced by distribution of \<^const>\<open>measurable\<close>-functions\<close>
-definition%important distr :: "'a measure \<Rightarrow> 'b measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b measure" where
+definition\<^marker>\<open>tag important\<close> distr :: "'a measure \<Rightarrow> 'b measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b measure" where
"distr M N f =
measure_of (space N) (sets N) (\<lambda>A. emeasure M (f -` A \<inter> space M))"
@@ -1519,7 +1519,7 @@
by (auto simp add: emeasure_distr measurable_space
intro!: arg_cong[where f="emeasure M"] measure_eqI)
-subsection%unimportant \<open>Real measure values\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Real measure values\<close>
lemma ring_of_finite_sets: "ring_of_sets (space M) {A\<in>sets M. emeasure M A \<noteq> top}"
proof (rule ring_of_setsI)
@@ -1709,7 +1709,7 @@
subsection \<open>Set of measurable sets with finite measure\<close>
-definition%important fmeasurable :: "'a measure \<Rightarrow> 'a set set" where
+definition\<^marker>\<open>tag important\<close> fmeasurable :: "'a measure \<Rightarrow> 'a set set" where
"fmeasurable M = {A\<in>sets M. emeasure M A < \<infinity>}"
lemma fmeasurableD[dest, measurable_dest]: "A \<in> fmeasurable M \<Longrightarrow> A \<in> sets M"
@@ -1749,7 +1749,7 @@
using emeasure_mono[of "a - b" a M] * by (auto simp: fmeasurable_def)
qed
-subsection%unimportant\<open>Measurable sets formed by unions and intersections\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Measurable sets formed by unions and intersections\<close>
lemma fmeasurable_Diff: "A \<in> fmeasurable M \<Longrightarrow> B \<in> sets M \<Longrightarrow> A - B \<in> fmeasurable M"
using fmeasurableI2[of A M "A - B"] by auto
@@ -2073,7 +2073,7 @@
subsection \<open>Measure spaces with \<^term>\<open>emeasure M (space M) < \<infinity>\<close>\<close>
-locale%important finite_measure = sigma_finite_measure M for M +
+locale\<^marker>\<open>tag important\<close> finite_measure = sigma_finite_measure M for M +
assumes finite_emeasure_space: "emeasure M (space M) \<noteq> top"
lemma finite_measureI[Pure.intro!]:
@@ -2295,7 +2295,7 @@
using bound[of x] sets_eq_imp_space_eq[OF sets] by (simp add: iter)
qed (auto simp add: iter le_fun_def INF_apply[abs_def] intro!: meas cont)
-subsection%unimportant \<open>Counting space\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Counting space\<close>
lemma strict_monoI_Suc:
assumes ord [simp]: "(\<And>n. f n < f (Suc n))" shows "strict_mono f"
@@ -2444,7 +2444,7 @@
show "sigma_finite_measure (count_space A)" ..
qed
-subsection%unimportant \<open>Measure restricted to space\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Measure restricted to space\<close>
lemma emeasure_restrict_space:
assumes "\<Omega> \<inter> space M \<in> sets M" "A \<subseteq> \<Omega>"
@@ -2603,7 +2603,7 @@
finally show "emeasure M X = emeasure N X" .
qed fact
-subsection%unimportant \<open>Null measure\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Null measure\<close>
definition null_measure :: "'a measure \<Rightarrow> 'a measure" where
"null_measure M = sigma (space M) (sets M)"
@@ -2627,7 +2627,7 @@
subsection \<open>Scaling a measure\<close>
-definition%important scale_measure :: "ennreal \<Rightarrow> 'a measure \<Rightarrow> 'a measure" where
+definition\<^marker>\<open>tag important\<close> scale_measure :: "ennreal \<Rightarrow> 'a measure \<Rightarrow> 'a measure" where
"scale_measure r M = measure_of (space M) (sets M) (\<lambda>A. r * emeasure M A)"
lemma space_scale_measure: "space (scale_measure r M) = space M"
@@ -2829,7 +2829,7 @@
done
qed
-text%important \<open>
+text\<^marker>\<open>tag important\<close> \<open>
Define a lexicographical order on \<^type>\<open>measure\<close>, in the order space, sets and measure. The parts
of the lexicographical order are point-wise ordered.
\<close>
@@ -2847,10 +2847,10 @@
if sets M = sets N then emeasure M \<le> emeasure N else sets M \<subseteq> sets N else space M \<subseteq> space N)"
by (auto elim: less_eq_measure.cases intro: less_eq_measure.intros)
-definition%important less_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> bool" where
+definition\<^marker>\<open>tag important\<close> less_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> bool" where
"less_measure M N \<longleftrightarrow> (M \<le> N \<and> \<not> N \<le> M)"
-definition%important bot_measure :: "'a measure" where
+definition\<^marker>\<open>tag important\<close> bot_measure :: "'a measure" where
"bot_measure = sigma {} {}"
lemma
@@ -2874,7 +2874,7 @@
by (cases "X \<in> sets M") (auto simp: emeasure_notin_sets)
done
-definition%important sup_measure' :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure" where
+definition\<^marker>\<open>tag important\<close> sup_measure' :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure" where
"sup_measure' A B =
measure_of (space A) (sets A)
(\<lambda>X. SUP Y\<in>sets A. emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y))"
@@ -3008,7 +3008,7 @@
qed
qed simp_all
-definition%important sup_lexord :: "'a \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow> 'b::order) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" where
+definition\<^marker>\<open>tag important\<close> sup_lexord :: "'a \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow> 'b::order) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" where
"sup_lexord A B k s c =
(if k A = k B then c else
if \<not> k A \<le> k B \<and> \<not> k B \<le> k A then s else
@@ -3038,7 +3038,7 @@
instantiation measure :: (type) semilattice_sup
begin
-definition%important sup_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure" where
+definition\<^marker>\<open>tag important\<close> sup_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure" where
"sup_measure A B =
sup_lexord A B space (sigma (space A \<union> space B) {})
(sup_lexord A B sets (sigma (space A) (sets A \<union> sets B)) (sup_measure' A B))"
@@ -3141,7 +3141,7 @@
lemma UN_space_closed: "\<Union>(sets ` S) \<subseteq> Pow (\<Union>(space ` S))"
using sets.space_closed by auto
-definition%important
+definition\<^marker>\<open>tag important\<close>
Sup_lexord :: "('a \<Rightarrow> 'b::complete_lattice) \<Rightarrow> ('a set \<Rightarrow> 'a) \<Rightarrow> ('a set \<Rightarrow> 'a) \<Rightarrow> 'a set \<Rightarrow> 'a"
where
"Sup_lexord k c s A =
@@ -3203,7 +3203,7 @@
"finite I \<Longrightarrow> I \<noteq> {} \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> sets i = sets M) \<Longrightarrow> sets (sup_measure.F id I) = sets M"
by (induction I rule: finite_ne_induct) (simp_all add: sets_sup)
-definition%important Sup_measure' :: "'a measure set \<Rightarrow> 'a measure" where
+definition\<^marker>\<open>tag important\<close> Sup_measure' :: "'a measure set \<Rightarrow> 'a measure" where
"Sup_measure' M =
measure_of (\<Union>a\<in>M. space a) (\<Union>a\<in>M. sets a)
(\<lambda>X. (SUP P\<in>{P. finite P \<and> P \<subseteq> M }. sup_measure.F id P X))"
@@ -3274,20 +3274,20 @@
using assms * by auto
qed (rule UN_space_closed)
-definition%important Sup_measure :: "'a measure set \<Rightarrow> 'a measure" where
+definition\<^marker>\<open>tag important\<close> Sup_measure :: "'a measure set \<Rightarrow> 'a measure" where
"Sup_measure =
Sup_lexord space
(Sup_lexord sets Sup_measure'
(\<lambda>U. sigma (\<Union>u\<in>U. space u) (\<Union>u\<in>U. sets u)))
(\<lambda>U. sigma (\<Union>u\<in>U. space u) {})"
-definition%important Inf_measure :: "'a measure set \<Rightarrow> 'a measure" where
+definition\<^marker>\<open>tag important\<close> Inf_measure :: "'a measure set \<Rightarrow> 'a measure" where
"Inf_measure A = Sup {x. \<forall>a\<in>A. x \<le> a}"
-definition%important inf_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure" where
+definition\<^marker>\<open>tag important\<close> inf_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure" where
"inf_measure a b = Inf {a, b}"
-definition%important top_measure :: "'a measure" where
+definition\<^marker>\<open>tag important\<close> top_measure :: "'a measure" where
"top_measure = Inf {}"
instance
@@ -3527,7 +3527,7 @@
qed
qed
-subsubsection%unimportant \<open>Supremum of a set of \<open>\<sigma>\<close>-algebras\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Supremum of a set of \<open>\<sigma>\<close>-algebras\<close>
lemma space_Sup_eq_UN: "space (Sup M) = (\<Union>x\<in>M. space x)"
unfolding Sup_measure_def
--- a/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy Fri Apr 12 22:09:25 2019 +0200
@@ -9,7 +9,7 @@
imports Measure_Space Borel_Space
begin
-subsection%unimportant \<open>Approximating functions\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Approximating functions\<close>
lemma AE_upper_bound_inf_ennreal:
fixes F G::"'a \<Rightarrow> ennreal"
@@ -115,7 +115,7 @@
\<close>
-definition%important "simple_function M g \<longleftrightarrow>
+definition\<^marker>\<open>tag important\<close> "simple_function M g \<longleftrightarrow>
finite (g ` space M) \<and>
(\<forall>x \<in> g ` space M. g -` {x} \<inter> space M \<in> sets M)"
@@ -301,11 +301,11 @@
shows "simple_function M (\<lambda>x. real (f x))"
by (rule simple_function_compose1[OF sf])
-lemma%important borel_measurable_implies_simple_function_sequence:
+lemma\<^marker>\<open>tag important\<close> borel_measurable_implies_simple_function_sequence:
fixes u :: "'a \<Rightarrow> ennreal"
assumes u[measurable]: "u \<in> borel_measurable M"
shows "\<exists>f. incseq f \<and> (\<forall>i. (\<forall>x. f i x < top) \<and> simple_function M (f i)) \<and> u = (SUP i. f i)"
-proof%unimportant -
+proof -
define f where [abs_def]:
"f i x = real_of_int (floor (enn2real (min i (u x)) * 2^i)) / 2^i" for i x
@@ -397,7 +397,7 @@
using borel_measurable_implies_simple_function_sequence [OF u]
by (metis SUP_apply)
-lemma%important simple_function_induct
+lemma\<^marker>\<open>tag important\<close> simple_function_induct
[consumes 1, case_names cong set mult add, induct set: simple_function]:
fixes u :: "'a \<Rightarrow> ennreal"
assumes u: "simple_function M u"
@@ -406,7 +406,7 @@
assumes mult: "\<And>u c. P u \<Longrightarrow> P (\<lambda>x. c * u x)"
assumes add: "\<And>u v. P u \<Longrightarrow> P v \<Longrightarrow> P (\<lambda>x. v x + u x)"
shows "P u"
-proof%unimportant (rule cong)
+proof (rule cong)
from AE_space show "AE x in M. (\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x) = u x"
proof eventually_elim
fix x assume x: "x \<in> space M"
@@ -468,7 +468,7 @@
qed fact
qed
-lemma%important borel_measurable_induct
+lemma\<^marker>\<open>tag important\<close> borel_measurable_induct
[consumes 1, case_names cong set mult add seq, induct set: borel_measurable]:
fixes u :: "'a \<Rightarrow> ennreal"
assumes u: "u \<in> borel_measurable M"
@@ -478,8 +478,8 @@
assumes add: "\<And>u v. u \<in> borel_measurable M\<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x < top) \<Longrightarrow> P u \<Longrightarrow> v \<in> borel_measurable M \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> v x < top) \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x = 0 \<or> v x = 0) \<Longrightarrow> P v \<Longrightarrow> P (\<lambda>x. v x + u x)"
assumes seq: "\<And>U. (\<And>i. U i \<in> borel_measurable M) \<Longrightarrow> (\<And>i x. x \<in> space M \<Longrightarrow> U i x < top) \<Longrightarrow> (\<And>i. P (U i)) \<Longrightarrow> incseq U \<Longrightarrow> u = (SUP i. U i) \<Longrightarrow> P (SUP i. U i)"
shows "P u"
- using%unimportant u
-proof%unimportant (induct rule: borel_measurable_implies_simple_function_sequence')
+ using u
+proof (induct rule: borel_measurable_implies_simple_function_sequence')
fix U assume U: "\<And>i. simple_function M (U i)" "incseq U" "\<And>i x. U i x < top" and sup: "\<And>x. (SUP i. U i x) = u x"
have u_eq: "u = (SUP i. U i)"
using u by (auto simp add: image_comp sup)
@@ -580,7 +580,7 @@
subsection "Simple integral"
-definition%important simple_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> ennreal" ("integral\<^sup>S") where
+definition\<^marker>\<open>tag important\<close> simple_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> ennreal" ("integral\<^sup>S") where
"integral\<^sup>S M f = (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M))"
syntax
@@ -813,7 +813,7 @@
subsection \<open>Integral on nonnegative functions\<close>
-definition%important nn_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> ennreal" ("integral\<^sup>N") where
+definition\<^marker>\<open>tag important\<close> nn_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> ennreal" ("integral\<^sup>N") where
"integral\<^sup>N M f = (SUP g \<in> {g. simple_function M g \<and> g \<le> f}. integral\<^sup>S M g)"
syntax
@@ -1645,7 +1645,7 @@
lemma nn_integral_bot[simp]: "nn_integral bot f = 0"
by (simp add: nn_integral_empty)
-subsubsection%unimportant \<open>Distributions\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Distributions\<close>
lemma nn_integral_distr:
assumes T: "T \<in> measurable M M'" and f: "f \<in> borel_measurable (distr M M' T)"
@@ -1670,7 +1670,7 @@
qed (simp_all add: measurable_compose[OF T] T nn_integral_cmult nn_integral_add
nn_integral_monotone_convergence_SUP le_fun_def incseq_def image_comp)
-subsubsection%unimportant \<open>Counting space\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Counting space\<close>
lemma simple_function_count_space[simp]:
"simple_function (count_space A) f \<longleftrightarrow> finite (f ` A)"
@@ -2083,7 +2083,7 @@
subsubsection \<open>Measure spaces with an associated density\<close>
-definition%important density :: "'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> 'a measure" where
+definition\<^marker>\<open>tag important\<close> density :: "'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> 'a measure" where
"density M f = measure_of (space M) (sets M) (\<lambda>A. \<integral>\<^sup>+ x. f x * indicator A x \<partial>M)"
lemma
@@ -2170,11 +2170,11 @@
by (intro exI[of _ "N \<union> {x\<in>space M. f x = 0}"] conjI *) (auto elim: eventually_elim2)
qed
-lemma%important nn_integral_density:
+lemma\<^marker>\<open>tag important\<close> nn_integral_density:
assumes f: "f \<in> borel_measurable M"
assumes g: "g \<in> borel_measurable M"
shows "integral\<^sup>N (density M f) g = (\<integral>\<^sup>+ x. f x * g x \<partial>M)"
-using%unimportant g proof%unimportant induct
+using g proof induct
case (cong u v)
then show ?case
apply (subst nn_integral_cong[OF cong(3)])
@@ -2293,7 +2293,7 @@
subsubsection \<open>Point measure\<close>
-definition%important point_measure :: "'a set \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> 'a measure" where
+definition\<^marker>\<open>tag important\<close> point_measure :: "'a set \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> 'a measure" where
"point_measure A f = density (count_space A) f"
lemma
@@ -2359,7 +2359,7 @@
subsubsection \<open>Uniform measure\<close>
-definition%important "uniform_measure M A = density M (\<lambda>x. indicator A x / emeasure M A)"
+definition\<^marker>\<open>tag important\<close> "uniform_measure M A = density M (\<lambda>x. indicator A x / emeasure M A)"
lemma
shows sets_uniform_measure[simp, measurable_cong]: "sets (uniform_measure M A) = sets M"
@@ -2434,7 +2434,7 @@
unfolding uniform_measure_def by (simp add: AE_density)
qed
-subsubsection%unimportant \<open>Null measure\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Null measure\<close>
lemma null_measure_eq_density: "null_measure M = density M (\<lambda>_. 0)"
by (intro measure_eqI) (simp_all add: emeasure_density)
@@ -2451,7 +2451,7 @@
subsubsection \<open>Uniform count measure\<close>
-definition%important "uniform_count_measure A = point_measure A (\<lambda>x. 1 / card A)"
+definition\<^marker>\<open>tag important\<close> "uniform_count_measure A = point_measure A (\<lambda>x. 1 / card A)"
lemma
shows space_uniform_count_measure: "space (uniform_count_measure A) = A"
@@ -2480,7 +2480,7 @@
"UNIV = sets (uniform_count_measure UNIV) \<longleftrightarrow> True"
by(simp_all add: sets_uniform_count_measure)
-subsubsection%unimportant \<open>Scaled measure\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Scaled measure\<close>
lemma nn_integral_scale_measure:
assumes f: "f \<in> borel_measurable M"
--- a/src/HOL/Analysis/Norm_Arith.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Norm_Arith.thy Fri Apr 12 22:09:25 2019 +0200
@@ -131,13 +131,13 @@
ML_file \<open>normarith.ML\<close>
-method_setup%important norm = \<open>
+method_setup\<^marker>\<open>tag important\<close> norm = \<open>
Scan.succeed (SIMPLE_METHOD' o NormArith.norm_arith_tac)
\<close> "prove simple linear statements about vector norms"
text \<open>Hence more metric properties.\<close>
-text%important \<open>%whitespace\<close>
+text\<^marker>\<open>tag important\<close> \<open>%whitespace\<close>
proposition dist_triangle_add:
fixes x y x' y' :: "'a::real_normed_vector"
shows "dist (x + y) (x' + y') \<le> dist x x' + dist y y'"
--- a/src/HOL/Analysis/Operator_Norm.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Operator_Norm.thy Fri Apr 12 22:09:25 2019 +0200
@@ -11,8 +11,8 @@
text \<open>This formulation yields zero if \<open>'a\<close> is the trivial vector space.\<close>
-text%important \<open>%whitespace\<close>
-definition%important
+text\<^marker>\<open>tag important\<close> \<open>%whitespace\<close>
+definition\<^marker>\<open>tag important\<close>
onorm :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> real" where
"onorm f = (SUP x. norm (f x) / norm x)"
--- a/src/HOL/Analysis/Ordered_Euclidean_Space.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Ordered_Euclidean_Space.thy Fri Apr 12 22:09:25 2019 +0200
@@ -157,7 +157,7 @@
shows "i \<in> Basis \<longleftrightarrow> fst i = 0 \<and> snd i \<in> Basis \<or> snd i = 0 \<and> fst i \<in> Basis"
by (cases i) (auto simp: Basis_prod_def)
-instantiation%unimportant prod :: (abs, abs) abs
+instantiation\<^marker>\<open>tag unimportant\<close> prod :: (abs, abs) abs
begin
definition "\<bar>x\<bar> = (\<bar>fst x\<bar>, \<bar>snd x\<bar>)"
@@ -299,11 +299,11 @@
instantiation vec :: (ordered_euclidean_space, finite) ordered_euclidean_space
begin
-definition%important "inf x y = (\<chi> i. inf (x $ i) (y $ i))"
-definition%important "sup x y = (\<chi> i. sup (x $ i) (y $ i))"
-definition%important "Inf X = (\<chi> i. (INF x\<in>X. x $ i))"
-definition%important "Sup X = (\<chi> i. (SUP x\<in>X. x $ i))"
-definition%important "\<bar>x\<bar> = (\<chi> i. \<bar>x $ i\<bar>)"
+definition\<^marker>\<open>tag important\<close> "inf x y = (\<chi> i. inf (x $ i) (y $ i))"
+definition\<^marker>\<open>tag important\<close> "sup x y = (\<chi> i. sup (x $ i) (y $ i))"
+definition\<^marker>\<open>tag important\<close> "Inf X = (\<chi> i. (INF x\<in>X. x $ i))"
+definition\<^marker>\<open>tag important\<close> "Sup X = (\<chi> i. (SUP x\<in>X. x $ i))"
+definition\<^marker>\<open>tag important\<close> "\<bar>x\<bar> = (\<chi> i. \<bar>x $ i\<bar>)"
instance
apply standard
--- a/src/HOL/Analysis/Path_Connected.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Path_Connected.thy Fri Apr 12 22:09:25 2019 +0200
@@ -10,34 +10,34 @@
subsection \<open>Paths and Arcs\<close>
-definition%important path :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> path :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> bool"
where "path g \<longleftrightarrow> continuous_on {0..1} g"
-definition%important pathstart :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> 'a"
+definition\<^marker>\<open>tag important\<close> pathstart :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> 'a"
where "pathstart g = g 0"
-definition%important pathfinish :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> 'a"
+definition\<^marker>\<open>tag important\<close> pathfinish :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> 'a"
where "pathfinish g = g 1"
-definition%important path_image :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> 'a set"
+definition\<^marker>\<open>tag important\<close> path_image :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> 'a set"
where "path_image g = g ` {0 .. 1}"
-definition%important reversepath :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> real \<Rightarrow> 'a"
+definition\<^marker>\<open>tag important\<close> reversepath :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> real \<Rightarrow> 'a"
where "reversepath g = (\<lambda>x. g(1 - x))"
-definition%important joinpaths :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> real \<Rightarrow> 'a"
+definition\<^marker>\<open>tag important\<close> joinpaths :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> real \<Rightarrow> 'a"
(infixr "+++" 75)
where "g1 +++ g2 = (\<lambda>x. if x \<le> 1/2 then g1 (2 * x) else g2 (2 * x - 1))"
-definition%important simple_path :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> simple_path :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> bool"
where "simple_path g \<longleftrightarrow>
path g \<and> (\<forall>x\<in>{0..1}. \<forall>y\<in>{0..1}. g x = g y \<longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0)"
-definition%important arc :: "(real \<Rightarrow> 'a :: topological_space) \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> arc :: "(real \<Rightarrow> 'a :: topological_space) \<Rightarrow> bool"
where "arc g \<longleftrightarrow> path g \<and> inj_on g {0..1}"
-subsection%unimportant\<open>Invariance theorems\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Invariance theorems\<close>
lemma path_eq: "path p \<Longrightarrow> (\<And>t. t \<in> {0..1} \<Longrightarrow> p t = q t) \<Longrightarrow> path q"
using continuous_on_eq path_def by blast
@@ -133,7 +133,7 @@
by (auto simp: arc_def inj_on_def path_linear_image_eq)
-subsection%unimportant\<open>Basic lemmas about paths\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Basic lemmas about paths\<close>
lemma pathin_iff_path_real [simp]: "pathin euclideanreal g \<longleftrightarrow> path g"
by (simp add: pathin_def path_def)
@@ -298,7 +298,7 @@
qed
-subsection%unimportant \<open>Path Images\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Path Images\<close>
lemma bounded_path_image: "path g \<Longrightarrow> bounded(path_image g)"
by (simp add: compact_imp_bounded compact_path_image)
@@ -391,7 +391,7 @@
by (auto simp: simple_path_def path_image_def inj_on_def less_eq_real_def Ball_def)
-subsection%unimportant\<open>Simple paths with the endpoints removed\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Simple paths with the endpoints removed\<close>
lemma simple_path_endless:
"simple_path c \<Longrightarrow> path_image c - {pathstart c,pathfinish c} = c ` {0<..<1}"
@@ -414,7 +414,7 @@
by (simp add: simple_path_endless)
-subsection%unimportant\<open>The operations on paths\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>The operations on paths\<close>
lemma path_image_subset_reversepath: "path_image(reversepath g) \<le> path_image g"
by (auto simp: path_image_def reversepath_def)
@@ -562,7 +562,7 @@
by (rule ext) (auto simp: mult.commute)
-subsection%unimportant\<open>Some reversed and "if and only if" versions of joining theorems\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Some reversed and "if and only if" versions of joining theorems\<close>
lemma path_join_path_ends:
fixes g1 :: "real \<Rightarrow> 'a::metric_space"
@@ -695,7 +695,7 @@
using pathfinish_in_path_image by (fastforce simp: arc_join_eq)
-subsection%unimportant\<open>The joining of paths is associative\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>The joining of paths is associative\<close>
lemma path_assoc:
"\<lbrakk>pathfinish p = pathstart q; pathfinish q = pathstart r\<rbrakk>
@@ -745,7 +745,7 @@
\<Longrightarrow> arc(p +++ (q +++ r)) \<longleftrightarrow> arc((p +++ q) +++ r)"
by (simp add: arc_simple_path simple_path_assoc)
-subsubsection%unimportant\<open>Symmetry and loops\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close>\<open>Symmetry and loops\<close>
lemma path_sym:
"\<lbrakk>pathfinish p = pathstart q; pathfinish q = pathstart p\<rbrakk> \<Longrightarrow> path(p +++ q) \<longleftrightarrow> path(q +++ p)"
@@ -764,7 +764,7 @@
subsection\<open>Subpath\<close>
-definition%important subpath :: "real \<Rightarrow> real \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> real \<Rightarrow> 'a::real_normed_vector"
+definition\<^marker>\<open>tag important\<close> subpath :: "real \<Rightarrow> real \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> real \<Rightarrow> 'a::real_normed_vector"
where "subpath a b g \<equiv> \<lambda>x. g((b - a) * x + a)"
lemma path_image_subpath_gen:
@@ -940,7 +940,7 @@
by (rule ext) (simp add: joinpaths_def subpath_def divide_simps)
-subsection%unimportant\<open>There is a subpath to the frontier\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>There is a subpath to the frontier\<close>
lemma subpath_to_frontier_explicit:
fixes S :: "'a::metric_space set"
@@ -1059,7 +1059,7 @@
subsection \<open>Shift Path to Start at Some Given Point\<close>
-definition%important shiftpath :: "real \<Rightarrow> (real \<Rightarrow> 'a::topological_space) \<Rightarrow> real \<Rightarrow> 'a"
+definition\<^marker>\<open>tag important\<close> shiftpath :: "real \<Rightarrow> (real \<Rightarrow> 'a::topological_space) \<Rightarrow> real \<Rightarrow> 'a"
where "shiftpath a f = (\<lambda>x. if (a + x) \<le> 1 then f (a + x) else f (a + x - 1))"
lemma pathstart_shiftpath: "a \<le> 1 \<Longrightarrow> pathstart (shiftpath a g) = g a"
@@ -1171,7 +1171,7 @@
subsection \<open>Straight-Line Paths\<close>
-definition%important linepath :: "'a::real_normed_vector \<Rightarrow> 'a \<Rightarrow> real \<Rightarrow> 'a"
+definition\<^marker>\<open>tag important\<close> linepath :: "'a::real_normed_vector \<Rightarrow> 'a \<Rightarrow> real \<Rightarrow> 'a"
where "linepath a b = (\<lambda>x. (1 - x) *\<^sub>R a + x *\<^sub>R b)"
lemma pathstart_linepath[simp]: "pathstart (linepath a b) = a"
@@ -1273,7 +1273,7 @@
using mult_left_le [of a "1-u"] mult_left_le [of b u] by auto
-subsection%unimportant\<open>Segments via convex hulls\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Segments via convex hulls\<close>
lemma segments_subset_convex_hull:
"closed_segment a b \<subseteq> (convex hull {a,b,c})"
@@ -1391,7 +1391,7 @@
qed
-subsection%unimportant \<open>Bounding a point away from a path\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Bounding a point away from a path\<close>
lemma not_on_path_ball:
fixes g :: "real \<Rightarrow> 'a::heine_borel"
@@ -1425,10 +1425,10 @@
text \<open>Original formalization by Tom Hales\<close>
-definition%important "path_component s x y \<longleftrightarrow>
+definition\<^marker>\<open>tag important\<close> "path_component s x y \<longleftrightarrow>
(\<exists>g. path g \<and> path_image g \<subseteq> s \<and> pathstart g = x \<and> pathfinish g = y)"
-abbreviation%important
+abbreviation\<^marker>\<open>tag important\<close>
"path_component_set s x \<equiv> Collect (path_component s x)"
lemmas path_defs = path_def pathstart_def pathfinish_def path_image_def path_component_def
@@ -1477,7 +1477,7 @@
unfolding path_component_def
by (rule_tac x="linepath a b" in exI, auto)
-subsubsection%unimportant \<open>Path components as sets\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Path components as sets\<close>
lemma path_component_set:
"path_component_set s x =
@@ -1502,7 +1502,7 @@
subsection \<open>Path connectedness of a space\<close>
-definition%important "path_connected s \<longleftrightarrow>
+definition\<^marker>\<open>tag important\<close> "path_connected s \<longleftrightarrow>
(\<forall>x\<in>s. \<forall>y\<in>s. \<exists>g. path g \<and> path_image g \<subseteq> s \<and> pathstart g = x \<and> pathfinish g = y)"
lemma path_connectedin_iff_path_connected_real [simp]:
@@ -1814,7 +1814,7 @@
qed
-subsection%unimportant\<open>Lemmas about path-connectedness\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Lemmas about path-connectedness\<close>
lemma path_connected_linear_image:
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
@@ -1909,7 +1909,7 @@
using is_interval_connected_1 is_interval_path_connected path_connected_imp_connected by blast
-subsection%unimportant\<open>Path components\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Path components\<close>
lemma Union_path_component [simp]:
"Union {path_component_set S x |x. x \<in> S} = S"
@@ -2736,7 +2736,7 @@
by (auto simp: path_connected_annulus [OF assms] path_connected_imp_connected)
-subsection%unimportant\<open>Relations between components and path components\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Relations between components and path components\<close>
lemma open_connected_component:
fixes s :: "'a::real_normed_vector set"
@@ -2841,7 +2841,7 @@
qed
-subsection%unimportant\<open>Existence of unbounded components\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Existence of unbounded components\<close>
lemma cobounded_unbounded_component:
fixes s :: "'a :: euclidean_space set"
@@ -2929,13 +2929,13 @@
subsection\<open>The \<open>inside\<close> and \<open>outside\<close> of a Set\<close>
-text%important\<open>The inside comprises the points in a bounded connected component of the set's complement.
+text\<^marker>\<open>tag important\<close>\<open>The inside comprises the points in a bounded connected component of the set's complement.
The outside comprises the points in unbounded connected component of the complement.\<close>
-definition%important inside where
+definition\<^marker>\<open>tag important\<close> inside where
"inside S \<equiv> {x. (x \<notin> S) \<and> bounded(connected_component_set ( - S) x)}"
-definition%important outside where
+definition\<^marker>\<open>tag important\<close> outside where
"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)}"
--- a/src/HOL/Analysis/Polytope.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Polytope.thy Fri Apr 12 22:09:25 2019 +0200
@@ -8,7 +8,7 @@
subsection \<open>Faces of a (usually convex) set\<close>
-definition%important face_of :: "['a::real_vector set, 'a set] \<Rightarrow> bool" (infixr "(face'_of)" 50)
+definition\<^marker>\<open>tag important\<close> face_of :: "['a::real_vector set, 'a set] \<Rightarrow> bool" (infixr "(face'_of)" 50)
where
"T face_of S \<longleftrightarrow>
T \<subseteq> S \<and> convex T \<and>
@@ -645,7 +645,7 @@
text\<open>That is, faces that are intersection with supporting hyperplane\<close>
-definition%important exposed_face_of :: "['a::euclidean_space set, 'a set] \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> exposed_face_of :: "['a::euclidean_space set, 'a set] \<Rightarrow> bool"
(infixr "(exposed'_face'_of)" 50)
where "T exposed_face_of S \<longleftrightarrow>
T face_of S \<and> (\<exists>a b. S \<subseteq> {x. a \<bullet> x \<le> b} \<and> T = S \<inter> {x. a \<bullet> x = b})"
@@ -876,7 +876,7 @@
subsection\<open>Extreme points of a set: its singleton faces\<close>
-definition%important extreme_point_of :: "['a::real_vector, 'a set] \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> extreme_point_of :: "['a::real_vector, 'a set] \<Rightarrow> bool"
(infixr "(extreme'_point'_of)" 50)
where "x extreme_point_of S \<longleftrightarrow>
x \<in> S \<and> (\<forall>a \<in> S. \<forall>b \<in> S. x \<notin> open_segment a b)"
@@ -976,7 +976,7 @@
subsection\<open>Facets\<close>
-definition%important facet_of :: "['a::euclidean_space set, 'a set] \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> facet_of :: "['a::euclidean_space set, 'a set] \<Rightarrow> bool"
(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"
@@ -1022,7 +1022,7 @@
subsection \<open>Edges: faces of affine dimension 1\<close> (*FIXME too small subsection, rearrange? *)
-definition%important edge_of :: "['a::euclidean_space set, 'a set] \<Rightarrow> bool" (infixr "(edge'_of)" 50)
+definition\<^marker>\<open>tag important\<close> edge_of :: "['a::euclidean_space set, 'a set] \<Rightarrow> bool" (infixr "(edge'_of)" 50)
where "e edge_of S \<longleftrightarrow> e face_of S \<and> aff_dim e = 1"
lemma edge_of_imp_subset:
@@ -1601,7 +1601,7 @@
subsection\<open>Polytopes\<close>
-definition%important polytope where
+definition\<^marker>\<open>tag important\<close> polytope where
"polytope S \<equiv> \<exists>v. finite v \<and> S = convex hull v"
lemma polytope_translation_eq: "polytope (image (\<lambda>x. a + x) S) \<longleftrightarrow> polytope S"
@@ -1696,7 +1696,7 @@
subsection\<open>Polyhedra\<close>
-definition%important polyhedron where
+definition\<^marker>\<open>tag important\<close> polyhedron where
"polyhedron S \<equiv>
\<exists>F. finite F \<and>
S = \<Inter> F \<and>
@@ -3088,7 +3088,7 @@
text\<open>The notion of n-simplex for integer \<^term>\<open>n \<ge> -1\<close>\<close>
-definition%important simplex :: "int \<Rightarrow> 'a::euclidean_space set \<Rightarrow> bool" (infix "simplex" 50)
+definition\<^marker>\<open>tag important\<close> simplex :: "int \<Rightarrow> 'a::euclidean_space set \<Rightarrow> bool" (infix "simplex" 50)
where "n simplex S \<equiv> \<exists>C. \<not> affine_dependent C \<and> int(card C) = n + 1 \<and> S = convex hull C"
lemma simplex:
@@ -3211,7 +3211,7 @@
subsection \<open>Simplicial complexes and triangulations\<close>
-definition%important simplicial_complex where
+definition\<^marker>\<open>tag important\<close> simplicial_complex where
"simplicial_complex \<C> \<equiv>
finite \<C> \<and>
(\<forall>S \<in> \<C>. \<exists>n. n simplex S) \<and>
@@ -3219,7 +3219,7 @@
(\<forall>S S'. S \<in> \<C> \<and> S' \<in> \<C>
\<longrightarrow> (S \<inter> S') face_of S \<and> (S \<inter> S') face_of S')"
-definition%important triangulation where
+definition\<^marker>\<open>tag important\<close> triangulation where
"triangulation \<T> \<equiv>
finite \<T> \<and>
(\<forall>T \<in> \<T>. \<exists>n. n simplex T) \<and>
--- a/src/HOL/Analysis/Product_Vector.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Product_Vector.thy Fri Apr 12 22:09:25 2019 +0200
@@ -23,10 +23,10 @@
definition scale :: "'a \<Rightarrow> 'b \<times> 'c \<Rightarrow> 'b \<times> 'c"
where "scale a v = (s1 a (fst v), s2 a (snd v))"
-lemma%important scale_prod: "scale x (a, b) = (s1 x a, s2 x b)"
+lemma\<^marker>\<open>tag important\<close> scale_prod: "scale x (a, b) = (s1 x a, s2 x b)"
by (auto simp: scale_def)
-sublocale%important p: module scale
+sublocale\<^marker>\<open>tag important\<close> p: module scale
proof qed (simp_all add: scale_def
m1.scale_left_distrib m1.scale_right_distrib m2.scale_left_distrib m2.scale_right_distrib)
@@ -113,7 +113,7 @@
(* TODO: Product of uniform spaces and compatibility with metric_spaces! *)
-instantiation%unimportant prod :: (metric_space, metric_space) dist
+instantiation\<^marker>\<open>tag unimportant\<close> prod :: (metric_space, metric_space) dist
begin
definition dist_prod_def[code del]:
@@ -122,7 +122,7 @@
instance ..
end
-instantiation%unimportant prod :: (metric_space, metric_space) uniformity_dist
+instantiation\<^marker>\<open>tag unimportant\<close> prod :: (metric_space, metric_space) uniformity_dist
begin
definition [code del]:
@@ -253,8 +253,8 @@
subsection \<open>Product is a Complete Metric Space\<close>
-instance%important prod :: (complete_space, complete_space) complete_space
-proof%unimportant
+instance\<^marker>\<open>tag important\<close> prod :: (complete_space, complete_space) complete_space
+proof
fix X :: "nat \<Rightarrow> 'a \<times> 'b" assume "Cauchy X"
have 1: "(\<lambda>n. fst (X n)) \<longlonglongrightarrow> lim (\<lambda>n. fst (X n))"
using Cauchy_fst [OF \<open>Cauchy X\<close>]
@@ -310,9 +310,9 @@
declare [[code abort: "norm::('a::real_normed_vector*'b::real_normed_vector) \<Rightarrow> real"]]
-instance%important prod :: (banach, banach) banach ..
+instance\<^marker>\<open>tag important\<close> prod :: (banach, banach) banach ..
-subsubsection%unimportant \<open>Pair operations are linear\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Pair operations are linear\<close>
lemma bounded_linear_fst: "bounded_linear fst"
using fst_add fst_scaleR
@@ -352,7 +352,7 @@
then show "\<exists>K. \<forall>x. norm (f x, g x) \<le> norm x * K" ..
qed
-subsubsection%unimportant \<open>Frechet derivatives involving pairs\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Frechet derivatives involving pairs\<close>
proposition has_derivative_Pair [derivative_intros]:
assumes f: "(f has_derivative f') (at x within s)"
@@ -387,7 +387,7 @@
unfolding split_beta' .
-subsubsection%unimportant \<open>Vector derivatives involving pairs\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Vector derivatives involving pairs\<close>
lemma has_vector_derivative_Pair[derivative_intros]:
assumes "(f has_vector_derivative f') (at x within s)"
--- a/src/HOL/Analysis/Radon_Nikodym.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Radon_Nikodym.thy Fri Apr 12 22:09:25 2019 +0200
@@ -8,7 +8,7 @@
imports Bochner_Integration
begin
-definition%important diff_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure"
+definition\<^marker>\<open>tag important\<close> diff_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure"
where
"diff_measure M N = measure_of (space M) (sets M) (\<lambda>A. emeasure M A - emeasure N A)"
@@ -165,7 +165,7 @@
subsection "Absolutely continuous"
-definition%important absolutely_continuous :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> bool" where
+definition\<^marker>\<open>tag important\<close> absolutely_continuous :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> bool" where
"absolutely_continuous M N \<longleftrightarrow> null_sets M \<subseteq> null_sets N"
lemma absolutely_continuousI_count_space: "absolutely_continuous (count_space A) M"
@@ -889,7 +889,7 @@
subsection \<open>Radon-Nikodym derivative\<close>
-definition%important RN_deriv :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a \<Rightarrow> ennreal" where
+definition\<^marker>\<open>tag important\<close> RN_deriv :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a \<Rightarrow> ennreal" where
"RN_deriv M N =
(if \<exists>f. f \<in> borel_measurable M \<and> density M f = N
then SOME f. f \<in> borel_measurable M \<and> density M f = N
--- a/src/HOL/Analysis/Riemann_Mapping.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Riemann_Mapping.thy Fri Apr 12 22:09:25 2019 +0200
@@ -10,7 +10,7 @@
subsection\<open>Moebius functions are biholomorphisms of the unit disc\<close>
-definition%important Moebius_function :: "[real,complex,complex] \<Rightarrow> complex" where
+definition\<^marker>\<open>tag important\<close> Moebius_function :: "[real,complex,complex] \<Rightarrow> complex" where
"Moebius_function \<equiv> \<lambda>t w z. exp(\<i> * of_real t) * (z - w) / (1 - cnj w * z)"
lemma Moebius_function_simple:
@@ -1411,7 +1411,7 @@
qed
-subsection%unimportant \<open>More Borsukian results\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>More Borsukian results\<close>
lemma Borsukian_componentwise_eq:
fixes S :: "'a::euclidean_space set"
--- a/src/HOL/Analysis/Set_Integral.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Set_Integral.thy Fri Apr 12 22:09:25 2019 +0200
@@ -15,11 +15,11 @@
Notation
*)
-definition%important "set_borel_measurable M A f \<equiv> (\<lambda>x. indicator A x *\<^sub>R f x) \<in> borel_measurable M"
+definition\<^marker>\<open>tag important\<close> "set_borel_measurable M A f \<equiv> (\<lambda>x. indicator A x *\<^sub>R f x) \<in> borel_measurable M"
-definition%important "set_integrable M A f \<equiv> integrable M (\<lambda>x. indicator A x *\<^sub>R f x)"
+definition\<^marker>\<open>tag important\<close> "set_integrable M A f \<equiv> integrable M (\<lambda>x. indicator A x *\<^sub>R f x)"
-definition%important "set_lebesgue_integral M A f \<equiv> lebesgue_integral M (\<lambda>x. indicator A x *\<^sub>R f x)"
+definition\<^marker>\<open>tag important\<close> "set_lebesgue_integral M A f \<equiv> lebesgue_integral M (\<lambda>x. indicator A x *\<^sub>R f x)"
syntax
"_ascii_set_lebesgue_integral" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real"
@@ -417,7 +417,7 @@
and intgbl: "set_integrable M (\<Union>i. A i) f"
shows "LINT x:(\<Union>i. A i)|M. f x = (\<Sum>i. (LINT x:(A i)|M. f x))"
unfolding set_lebesgue_integral_def
-proof%unimportant (subst integral_suminf[symmetric])
+proof (subst integral_suminf[symmetric])
show int_A: "integrable M (\<lambda>x. indicat_real (A i) x *\<^sub>R f x)" for i
using intgbl unfolding set_integrable_def [symmetric]
by (rule set_integrable_subset) auto
--- a/src/HOL/Analysis/Sigma_Algebra.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Sigma_Algebra.thy Fri Apr 12 22:09:25 2019 +0200
@@ -30,7 +30,7 @@
subsection \<open>Families of sets\<close>
-locale%important subset_class =
+locale\<^marker>\<open>tag important\<close> subset_class =
fixes \<Omega> :: "'a set" and M :: "'a set set"
assumes space_closed: "M \<subseteq> Pow \<Omega>"
@@ -39,7 +39,7 @@
subsubsection \<open>Semiring of sets\<close>
-locale%important semiring_of_sets = subset_class +
+locale\<^marker>\<open>tag important\<close> semiring_of_sets = subset_class +
assumes empty_sets[iff]: "{} \<in> M"
assumes Int[intro]: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<inter> b \<in> M"
assumes Diff_cover:
@@ -76,7 +76,7 @@
subsubsection \<open>Ring of sets\<close>
-locale%important ring_of_sets = semiring_of_sets +
+locale\<^marker>\<open>tag important\<close> ring_of_sets = semiring_of_sets +
assumes Un [intro]: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<union> b \<in> M"
lemma (in ring_of_sets) finite_Union [intro]:
@@ -142,7 +142,7 @@
subsubsection \<open>Algebra of sets\<close>
-locale%important algebra = ring_of_sets +
+locale\<^marker>\<open>tag important\<close> algebra = ring_of_sets +
assumes top [iff]: "\<Omega> \<in> M"
lemma (in algebra) compl_sets [intro]:
@@ -221,7 +221,7 @@
"X \<subseteq> S \<Longrightarrow> algebra S { {}, X, S - X, S }"
by (auto simp: algebra_iff_Int)
-subsubsection%unimportant \<open>Restricted algebras\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Restricted algebras\<close>
abbreviation (in algebra)
"restricted_space A \<equiv> ((\<inter>) A) ` M"
@@ -232,7 +232,7 @@
subsubsection \<open>Sigma Algebras\<close>
-locale%important sigma_algebra = algebra +
+locale\<^marker>\<open>tag important\<close> sigma_algebra = algebra +
assumes countable_nat_UN [intro]: "\<And>A. range A \<subseteq> M \<Longrightarrow> (\<Union>i::nat. A i) \<in> M"
lemma (in algebra) is_sigma_algebra:
@@ -431,7 +431,7 @@
shows "sigma_algebra S { {}, X, S - X, S }"
using algebra.is_sigma_algebra[OF algebra_single_set[OF \<open>X \<subseteq> S\<close>]] by simp
-subsubsection%unimportant \<open>Binary Unions\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Binary Unions\<close>
definition binary :: "'a \<Rightarrow> 'a \<Rightarrow> nat \<Rightarrow> 'a"
where "binary a b = (\<lambda>x. b)(0 := a)"
@@ -473,10 +473,10 @@
subsubsection \<open>Initial Sigma Algebra\<close>
-text%important \<open>Sigma algebras can naturally be created as the closure of any set of
+text\<^marker>\<open>tag important\<close> \<open>Sigma algebras can naturally be created as the closure of any set of
M with regard to the properties just postulated.\<close>
-inductive_set%important sigma_sets :: "'a set \<Rightarrow> 'a set set \<Rightarrow> 'a set set"
+inductive_set\<^marker>\<open>tag important\<close> sigma_sets :: "'a set \<Rightarrow> 'a set set \<Rightarrow> 'a set set"
for sp :: "'a set" and A :: "'a set set"
where
Basic[intro, simp]: "a \<in> A \<Longrightarrow> a \<in> sigma_sets sp A"
@@ -819,7 +819,7 @@
thus "(\<Union>i::nat. A i) \<in> M" by (simp add: UN_disjointed_eq)
qed
-subsubsection%unimportant \<open>Ring generated by a semiring\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Ring generated by a semiring\<close>
definition (in semiring_of_sets) generated_ring :: "'a set set" where
"generated_ring = { \<Union>C | C. C \<subseteq> M \<and> finite C \<and> disjoint C }"
@@ -950,7 +950,7 @@
by (blast intro!: sigma_sets_mono elim: generated_ringE)
qed (auto intro!: generated_ringI_Basic sigma_sets_mono)
-subsubsection%unimportant \<open>A Two-Element Series\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>A Two-Element Series\<close>
definition binaryset :: "'a set \<Rightarrow> 'a set \<Rightarrow> nat \<Rightarrow> 'a set"
where "binaryset A B = (\<lambda>x. {})(0 := A, Suc 0 := B)"
@@ -966,7 +966,7 @@
subsubsection \<open>Closed CDI\<close>
-definition%important closed_cdi :: "'a set \<Rightarrow> 'a set set \<Rightarrow> bool" where
+definition\<^marker>\<open>tag important\<close> closed_cdi :: "'a set \<Rightarrow> 'a set set \<Rightarrow> bool" where
"closed_cdi \<Omega> M \<longleftrightarrow>
M \<subseteq> Pow \<Omega> &
(\<forall>s \<in> M. \<Omega> - s \<in> M) &
@@ -1200,7 +1200,7 @@
subsubsection \<open>Dynkin systems\<close>
-locale%important Dynkin_system = subset_class +
+locale\<^marker>\<open>tag important\<close> Dynkin_system = subset_class +
assumes space: "\<Omega> \<in> M"
and compl[intro!]: "\<And>A. A \<in> M \<Longrightarrow> \<Omega> - A \<in> M"
and UN[intro!]: "\<And>A. disjoint_family A \<Longrightarrow> range A \<subseteq> M
@@ -1262,7 +1262,7 @@
subsubsection "Intersection sets systems"
-definition%important Int_stable :: "'a set set \<Rightarrow> bool" where
+definition\<^marker>\<open>tag important\<close> Int_stable :: "'a set set \<Rightarrow> bool" where
"Int_stable M \<longleftrightarrow> (\<forall> a \<in> M. \<forall> b \<in> M. a \<inter> b \<in> M)"
lemma (in algebra) Int_stable: "Int_stable M"
@@ -1303,7 +1303,7 @@
subsubsection "Smallest Dynkin systems"
-definition%important Dynkin :: "'a set \<Rightarrow> 'a set set \<Rightarrow> 'a set set" where
+definition\<^marker>\<open>tag important\<close> Dynkin :: "'a set \<Rightarrow> 'a set set \<Rightarrow> 'a set set" where
"Dynkin \<Omega> M = (\<Inter>{D. Dynkin_system \<Omega> D \<and> M \<subseteq> D})"
lemma Dynkin_system_Dynkin:
@@ -1450,7 +1450,7 @@
subsubsection \<open>Induction rule for intersection-stable generators\<close>
-text%important \<open>The reason to introduce Dynkin-systems is the following induction rules for \<open>\<sigma>\<close>-algebras
+text\<^marker>\<open>tag important\<close> \<open>The reason to introduce Dynkin-systems is the following induction rules for \<open>\<sigma>\<close>-algebras
generated by a generator closed under intersection.\<close>
proposition sigma_sets_induct_disjoint[consumes 3, case_names basic empty compl union]:
@@ -1476,37 +1476,37 @@
subsection \<open>Measure type\<close>
-definition%important positive :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> bool" where
+definition\<^marker>\<open>tag important\<close> positive :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> bool" where
"positive M \<mu> \<longleftrightarrow> \<mu> {} = 0"
-definition%important countably_additive :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> bool" where
+definition\<^marker>\<open>tag important\<close> countably_additive :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> bool" where
"countably_additive M f \<longleftrightarrow>
(\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow>
(\<Sum>i. f (A i)) = f (\<Union>i. A i))"
-definition%important measure_space :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> bool" where
+definition\<^marker>\<open>tag important\<close> measure_space :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> bool" where
"measure_space \<Omega> A \<mu> \<longleftrightarrow>
sigma_algebra \<Omega> A \<and> positive A \<mu> \<and> countably_additive A \<mu>"
-typedef%important 'a measure =
+typedef\<^marker>\<open>tag important\<close> 'a measure =
"{(\<Omega>::'a set, A, \<mu>). (\<forall>a\<in>-A. \<mu> a = 0) \<and> measure_space \<Omega> A \<mu> }"
-proof%unimportant
+proof
have "sigma_algebra UNIV {{}, UNIV}"
by (auto simp: sigma_algebra_iff2)
then show "(UNIV, {{}, UNIV}, \<lambda>A. 0) \<in> {(\<Omega>, A, \<mu>). (\<forall>a\<in>-A. \<mu> a = 0) \<and> measure_space \<Omega> A \<mu>} "
by (auto simp: measure_space_def positive_def countably_additive_def)
qed
-definition%important space :: "'a measure \<Rightarrow> 'a set" where
+definition\<^marker>\<open>tag important\<close> space :: "'a measure \<Rightarrow> 'a set" where
"space M = fst (Rep_measure M)"
-definition%important sets :: "'a measure \<Rightarrow> 'a set set" where
+definition\<^marker>\<open>tag important\<close> sets :: "'a measure \<Rightarrow> 'a set set" where
"sets M = fst (snd (Rep_measure M))"
-definition%important emeasure :: "'a measure \<Rightarrow> 'a set \<Rightarrow> ennreal" where
+definition\<^marker>\<open>tag important\<close> emeasure :: "'a measure \<Rightarrow> 'a set \<Rightarrow> ennreal" where
"emeasure M = snd (snd (Rep_measure M))"
-definition%important measure :: "'a measure \<Rightarrow> 'a set \<Rightarrow> real" where
+definition\<^marker>\<open>tag important\<close> measure :: "'a measure \<Rightarrow> 'a set \<Rightarrow> real" where
"measure M A = enn2real (emeasure M A)"
declare [[coercion sets]]
@@ -1521,7 +1521,7 @@
interpretation sets: sigma_algebra "space M" "sets M" for M :: "'a measure"
using measure_space[of M] by (auto simp: measure_space_def)
-definition%important measure_of :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> 'a measure"
+definition\<^marker>\<open>tag important\<close> measure_of :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> 'a measure"
where
"measure_of \<Omega> A \<mu> =
Abs_measure (\<Omega>, if A \<subseteq> Pow \<Omega> then sigma_sets \<Omega> A else {{}, \<Omega>},
@@ -1740,7 +1740,7 @@
subsubsection \<open>Measurable functions\<close>
-definition%important measurable :: "'a measure \<Rightarrow> 'b measure \<Rightarrow> ('a \<Rightarrow> 'b) set"
+definition\<^marker>\<open>tag important\<close> measurable :: "'a measure \<Rightarrow> 'b measure \<Rightarrow> ('a \<Rightarrow> 'b) set"
(infixr "\<rightarrow>\<^sub>M" 60) where
"measurable A B = {f \<in> space A \<rightarrow> space B. \<forall>y \<in> sets B. f -` y \<inter> space A \<in> sets A}"
@@ -1903,7 +1903,7 @@
subsubsection \<open>Counting space\<close>
-definition%important count_space :: "'a set \<Rightarrow> 'a measure" where
+definition\<^marker>\<open>tag important\<close> count_space :: "'a set \<Rightarrow> 'a measure" where
"count_space \<Omega> = measure_of \<Omega> (Pow \<Omega>) (\<lambda>A. if finite A then of_nat (card A) else \<infinity>)"
lemma
@@ -1979,7 +1979,7 @@
"space N = {} \<Longrightarrow> f \<in> measurable M N \<longleftrightarrow> space M = {}"
by (auto simp add: measurable_def Pi_iff)
-subsubsection%unimportant \<open>Extend measure\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Extend measure\<close>
definition extend_measure :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('b \<Rightarrow> 'a set) \<Rightarrow> ('b \<Rightarrow> ennreal) \<Rightarrow> 'a measure"
where
@@ -2043,7 +2043,7 @@
subsection \<open>The smallest \<open>\<sigma>\<close>-algebra regarding a function\<close>
-definition%important vimage_algebra :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b measure \<Rightarrow> 'a measure" where
+definition\<^marker>\<open>tag important\<close> vimage_algebra :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b measure \<Rightarrow> 'a measure" where
"vimage_algebra X f M = sigma X {f -` A \<inter> X | A. A \<in> sets M}"
lemma space_vimage_algebra[simp]: "space (vimage_algebra X f M) = X"
--- a/src/HOL/Analysis/Starlike.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Starlike.thy Fri Apr 12 22:09:25 2019 +0200
@@ -15,7 +15,7 @@
subsection \<open>Midpoint\<close>
-definition%important midpoint :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a"
+definition\<^marker>\<open>tag important\<close> midpoint :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a"
where "midpoint a b = (inverse (2::real)) *\<^sub>R (a + b)"
lemma midpoint_idem [simp]: "midpoint x x = x"
@@ -91,10 +91,10 @@
subsection \<open>Line segments\<close>
-definition%important closed_segment :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a set"
+definition\<^marker>\<open>tag important\<close> closed_segment :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a set"
where "closed_segment a b = {(1 - u) *\<^sub>R a + u *\<^sub>R b | u::real. 0 \<le> u \<and> u \<le> 1}"
-definition%important open_segment :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a set" where
+definition\<^marker>\<open>tag important\<close> open_segment :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a set" where
"open_segment a b \<equiv> closed_segment a b - {a,b}"
lemmas segment = open_segment_def closed_segment_def
@@ -618,7 +618,7 @@
subsection\<open>Starlike sets\<close>
-definition%important "starlike S \<longleftrightarrow> (\<exists>a\<in>S. \<forall>x\<in>S. closed_segment a x \<subseteq> S)"
+definition\<^marker>\<open>tag important\<close> "starlike S \<longleftrightarrow> (\<exists>a\<in>S. \<forall>x\<in>S. closed_segment a x \<subseteq> S)"
lemma starlike_UNIV [simp]: "starlike UNIV"
by (simp add: starlike_def)
@@ -681,7 +681,7 @@
by (meson hull_mono inf_mono subset_insertI subset_refl)
qed
-subsection%unimportant\<open>More results about segments\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>More results about segments\<close>
lemma dist_half_times2:
fixes a :: "'a :: real_normed_vector"
@@ -903,7 +903,7 @@
subsection\<open>Betweenness\<close>
-definition%important "between = (\<lambda>(a,b) x. x \<in> closed_segment a b)"
+definition\<^marker>\<open>tag important\<close> "between = (\<lambda>(a,b) x. x \<in> closed_segment a b)"
lemma betweenI:
assumes "0 \<le> u" "u \<le> 1" "x = (1 - u) *\<^sub>R a + u *\<^sub>R b"
@@ -1059,7 +1059,7 @@
by (auto simp: between_mem_segment closed_segment_eq_real_ivl)
-subsection%unimportant \<open>Shrinking towards the interior of a convex set\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Shrinking towards the interior of a convex set\<close>
lemma mem_interior_convex_shrink:
fixes S :: "'a::euclidean_space set"
@@ -1247,7 +1247,7 @@
qed
-subsection%unimportant \<open>Some obvious but surprisingly hard simplex lemmas\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Some obvious but surprisingly hard simplex lemmas\<close>
lemma simplex:
assumes "finite S"
@@ -1631,7 +1631,7 @@
qed
-subsection%unimportant \<open>Relative interior of convex set\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Relative interior of convex set\<close>
lemma rel_interior_convex_nonempty_aux:
fixes S :: "'n::euclidean_space set"
@@ -2019,7 +2019,7 @@
subsection\<open>The relative frontier of a set\<close>
-definition%important "rel_frontier S = closure S - rel_interior S"
+definition\<^marker>\<open>tag important\<close> "rel_frontier S = closure S - rel_interior S"
lemma rel_frontier_empty [simp]: "rel_frontier {} = {}"
by (simp add: rel_frontier_def)
@@ -2393,7 +2393,7 @@
qed
-subsubsection%unimportant \<open>Relative interior and closure under common operations\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Relative interior and closure under common operations\<close>
lemma rel_interior_inter_aux: "\<Inter>{rel_interior S |S. S \<in> I} \<subseteq> \<Inter>I"
proof -
@@ -3085,7 +3085,7 @@
by (force simp: rel_frontier_def rel_interior_Times assms closure_Times)
-subsubsection%unimportant \<open>Relative interior of convex cone\<close>
+subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Relative interior of convex cone\<close>
lemma cone_rel_interior:
fixes S :: "'m::euclidean_space set"
@@ -3358,7 +3358,7 @@
qed
-subsection%unimportant \<open>Convexity on direct sums\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Convexity on direct sums\<close>
lemma closure_sum:
fixes S T :: "'a::real_normed_vector set"
@@ -3723,7 +3723,7 @@
using \<open>y < x\<close> by (simp add: field_simps)
qed simp
-subsection%unimportant\<open>Explicit formulas for interior and relative interior of convex hull\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Explicit formulas for interior and relative interior of convex hull\<close>
lemma at_within_cbox_finite:
assumes "x \<in> box a b" "x \<notin> S" "finite S"
@@ -4084,7 +4084,7 @@
by (metis Diff_cancel convex_hull_singleton insert_absorb2 open_segment_def segment_convex_hull)
qed
-subsection%unimportant\<open>Similar results for closure and (relative or absolute) frontier\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Similar results for closure and (relative or absolute) frontier\<close>
lemma closure_convex_hull [simp]:
fixes s :: "'a::euclidean_space set"
@@ -4242,7 +4242,7 @@
subsection \<open>Coplanarity, and collinearity in terms of affine hull\<close>
-definition%important coplanar where
+definition\<^marker>\<open>tag important\<close> coplanar where
"coplanar s \<equiv> \<exists>u v w. s \<subseteq> affine hull {u,v,w}"
lemma collinear_affine_hull:
@@ -4693,7 +4693,7 @@
done
-subsection%unimportant\<open>Basic lemmas about hyperplanes and halfspaces\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Basic lemmas about hyperplanes and halfspaces\<close>
lemma halfspace_Int_eq:
"{x. a \<bullet> x \<le> b} \<inter> {x. b \<le> a \<bullet> x} = {x. a \<bullet> x = b}"
@@ -4751,9 +4751,9 @@
using halfspace_eq_empty_le [of "-a" "-b"]
by simp
-subsection%unimportant\<open>Use set distance for an easy proof of separation properties\<close>
-
-proposition%unimportant separation_closures:
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Use set distance for an easy proof of separation properties\<close>
+
+proposition\<^marker>\<open>tag unimportant\<close> separation_closures:
fixes S :: "'a::euclidean_space set"
assumes "S \<inter> closure T = {}" "T \<inter> closure S = {}"
obtains U V where "U \<inter> V = {}" "open U" "open V" "S \<subseteq> U" "T \<subseteq> V"
@@ -5079,7 +5079,7 @@
by (simp add: continuous_on_closed * closedin_imp_subset)
qed
-subsection%unimportant\<open>Trivial fact: convexity equals connectedness for collinear sets\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Trivial fact: convexity equals connectedness for collinear sets\<close>
lemma convex_connected_collinear:
fixes S :: "'a::euclidean_space set"
@@ -5297,9 +5297,9 @@
by (simp add: clo closedin_closed_Int)
qed
-subsubsection%unimportant\<open>Representing affine hull as a finite intersection of hyperplanes\<close>
-
-proposition%unimportant affine_hull_convex_Int_nonempty_interior:
+subsubsection\<^marker>\<open>tag unimportant\<close>\<open>Representing affine hull as a finite intersection of hyperplanes\<close>
+
+proposition\<^marker>\<open>tag unimportant\<close> affine_hull_convex_Int_nonempty_interior:
fixes S :: "'a::real_normed_vector set"
assumes "convex S" "S \<inter> interior T \<noteq> {}"
shows "affine hull (S \<inter> T) = affine hull S"
@@ -5524,7 +5524,7 @@
shows "a \<noteq> 0 \<Longrightarrow> aff_dim {x. a \<bullet> x = r} = DIM('a) - 1"
by (metis aff_dim_eq_hyperplane affine_hull_eq affine_hyperplane)
-subsection%unimportant\<open>Some stepping theorems\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Some stepping theorems\<close>
lemma aff_dim_insert:
fixes a :: "'a::euclidean_space"
@@ -5649,9 +5649,9 @@
by (metis One_nat_def \<open>a \<noteq> 0\<close> affine_bounded_eq_lowdim affine_hyperplane bounded_hyperplane_eq_trivial_0)
qed
-subsection%unimportant\<open>General case without assuming closure and getting non-strict separation\<close>
-
-proposition%unimportant separating_hyperplane_closed_point_inset:
+subsection\<^marker>\<open>tag unimportant\<close>\<open>General case without assuming closure and getting non-strict separation\<close>
+
+proposition\<^marker>\<open>tag unimportant\<close> separating_hyperplane_closed_point_inset:
fixes S :: "'a::euclidean_space set"
assumes "convex S" "closed S" "S \<noteq> {}" "z \<notin> S"
obtains a b where "a \<in> S" "(a - z) \<bullet> z < b" "\<And>x. x \<in> S \<Longrightarrow> b < (a - z) \<bullet> x"
@@ -5688,7 +5688,7 @@
by simp (metis \<open>0 \<notin> S\<close>)
-proposition%unimportant separating_hyperplane_set_0_inspan:
+proposition\<^marker>\<open>tag unimportant\<close> separating_hyperplane_set_0_inspan:
fixes S :: "'a::euclidean_space set"
assumes "convex S" "S \<noteq> {}" "0 \<notin> S"
obtains a where "a \<in> span S" "a \<noteq> 0" "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> a \<bullet> x"
@@ -5779,7 +5779,7 @@
done
qed
-proposition%unimportant supporting_hyperplane_rel_boundary:
+proposition\<^marker>\<open>tag unimportant\<close> supporting_hyperplane_rel_boundary:
fixes S :: "'a::euclidean_space set"
assumes "convex S" "x \<in> S" and xno: "x \<notin> rel_interior S"
obtains a where "a \<noteq> 0"
@@ -5837,7 +5837,7 @@
by (metis assms convex_closure convex_rel_interior_closure)
-subsection%unimportant\<open> Some results on decomposing convex hulls: intersections, simplicial subdivision\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open> Some results on decomposing convex hulls: intersections, simplicial subdivision\<close>
lemma
fixes s :: "'a::euclidean_space set"
@@ -5882,7 +5882,7 @@
qed
-proposition%unimportant affine_hull_Int:
+proposition\<^marker>\<open>tag unimportant\<close> affine_hull_Int:
fixes s :: "'a::euclidean_space set"
assumes "\<not> affine_dependent(s \<union> t)"
shows "affine hull (s \<inter> t) = affine hull s \<inter> affine hull t"
@@ -5890,7 +5890,7 @@
apply (simp add: hull_mono)
by (simp add: affine_hull_Int_subset assms)
-proposition%unimportant convex_hull_Int:
+proposition\<^marker>\<open>tag unimportant\<close> convex_hull_Int:
fixes s :: "'a::euclidean_space set"
assumes "\<not> affine_dependent(s \<union> t)"
shows "convex hull (s \<inter> t) = convex hull s \<inter> convex hull t"
@@ -5898,7 +5898,7 @@
apply (simp add: hull_mono)
by (simp add: convex_hull_Int_subset assms)
-proposition%unimportant
+proposition\<^marker>\<open>tag unimportant\<close>
fixes s :: "'a::euclidean_space set set"
assumes "\<not> affine_dependent (\<Union>s)"
shows affine_hull_Inter: "affine hull (\<Inter>s) = (\<Inter>t\<in>s. affine hull t)" (is "?A")
@@ -5928,7 +5928,7 @@
by auto
qed
-proposition%unimportant in_convex_hull_exchange_unique:
+proposition\<^marker>\<open>tag unimportant\<close> in_convex_hull_exchange_unique:
fixes S :: "'a::euclidean_space set"
assumes naff: "\<not> affine_dependent S" and a: "a \<in> convex hull S"
and S: "T \<subseteq> S" "T' \<subseteq> S"
@@ -6076,7 +6076,7 @@
qed
qed
-corollary%unimportant convex_hull_exchange_Int:
+corollary\<^marker>\<open>tag unimportant\<close> convex_hull_exchange_Int:
fixes a :: "'a::euclidean_space"
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')) =
@@ -6229,7 +6229,7 @@
by (simp add: subs) (metis (lifting) span_eq_iff subs)
qed
-proposition%unimportant affine_hyperplane_sums_eq_UNIV:
+proposition\<^marker>\<open>tag unimportant\<close> affine_hyperplane_sums_eq_UNIV:
fixes S :: "'a :: euclidean_space set"
assumes "affine S"
and "S \<inter> {v. a \<bullet> v = b} \<noteq> {}"
@@ -6466,7 +6466,7 @@
by force
qed
-corollary%unimportant dense_complement_affine:
+corollary\<^marker>\<open>tag unimportant\<close> dense_complement_affine:
fixes S :: "'a :: euclidean_space set"
assumes less: "aff_dim T < aff_dim S" and "affine S" shows "closure(S - T) = S"
proof (cases "S \<inter> T = {}")
@@ -6487,7 +6487,7 @@
by (metis closure_translation translation_diff translation_invert)
qed
-corollary%unimportant dense_complement_openin_affine_hull:
+corollary\<^marker>\<open>tag unimportant\<close> dense_complement_openin_affine_hull:
fixes S :: "'a :: euclidean_space set"
assumes less: "aff_dim T < aff_dim S"
and ope: "openin (top_of_set (affine hull S)) S"
@@ -6501,7 +6501,7 @@
by (metis Int_Diff aff_dim_affine_hull affine_affine_hull dense_complement_affine hull_subset inf.orderE less)
qed
-corollary%unimportant dense_complement_convex:
+corollary\<^marker>\<open>tag unimportant\<close> dense_complement_convex:
fixes S :: "'a :: euclidean_space set"
assumes "aff_dim T < aff_dim S" "convex S"
shows "closure(S - T) = closure S"
@@ -6516,19 +6516,19 @@
by (metis Diff_mono \<open>convex S\<close> closure_mono convex_closure_rel_interior order_refl rel_interior_subset)
qed
-corollary%unimportant dense_complement_convex_closed:
+corollary\<^marker>\<open>tag unimportant\<close> dense_complement_convex_closed:
fixes S :: "'a :: euclidean_space set"
assumes "aff_dim T < aff_dim S" "convex S" "closed S"
shows "closure(S - T) = S"
by (simp add: assms dense_complement_convex)
-subsection%unimportant\<open>Parallel slices, etc\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Parallel slices, etc\<close>
text\<open> If we take a slice out of a set, we can do it perpendicularly,
with the normal vector to the slice parallel to the affine hull.\<close>
-proposition%unimportant affine_parallel_slice:
+proposition\<^marker>\<open>tag unimportant\<close> affine_parallel_slice:
fixes S :: "'a :: euclidean_space set"
assumes "affine S"
and "S \<inter> {x. a \<bullet> x \<le> b} \<noteq> {}"
@@ -6845,7 +6845,7 @@
qed
qed
-corollary%unimportant paracompact_closed:
+corollary\<^marker>\<open>tag unimportant\<close> paracompact_closed:
fixes S :: "'a :: {metric_space,second_countable_topology} set"
assumes "closed S"
and opC: "\<And>T. T \<in> \<C> \<Longrightarrow> open T"
@@ -6857,7 +6857,7 @@
by (rule paracompact_closedin [of UNIV S \<C>]) (auto simp: assms)
-subsection%unimportant\<open>Closed-graph characterization of continuity\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Closed-graph characterization of continuity\<close>
lemma continuous_closed_graph_gen:
fixes T :: "'b::real_normed_vector set"
@@ -6971,9 +6971,9 @@
by simp
qed
-subsection%unimportant\<open>The union of two collinear segments is another segment\<close>
-
-proposition%unimportant in_convex_hull_exchange:
+subsection\<^marker>\<open>tag unimportant\<close>\<open>The union of two collinear segments is another segment\<close>
+
+proposition\<^marker>\<open>tag unimportant\<close> in_convex_hull_exchange:
fixes a :: "'a::euclidean_space"
assumes a: "a \<in> convex hull S" and xS: "x \<in> convex hull S"
obtains b where "b \<in> S" "x \<in> convex hull (insert a (S - {b}))"
@@ -7250,7 +7250,7 @@
subsection\<open>Orthogonal complement\<close>
-definition%important orthogonal_comp ("_\<^sup>\<bottom>" [80] 80)
+definition\<^marker>\<open>tag important\<close> orthogonal_comp ("_\<^sup>\<bottom>" [80] 80)
where "orthogonal_comp W \<equiv> {x. \<forall>y \<in> W. orthogonal y x}"
proposition subspace_orthogonal_comp: "subspace (W\<^sup>\<bottom>)"
@@ -7366,7 +7366,7 @@
apply (simp add: adjoint_works assms(1) inner_commute)
by (metis adjoint_works all_zero_iff assms(1) inner_commute)
-subsection%unimportant \<open>A non-injective linear function maps into a hyperplane.\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>A non-injective linear function maps into a hyperplane.\<close>
lemma linear_surj_adj_imp_inj:
fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
--- a/src/HOL/Analysis/Summation_Tests.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Summation_Tests.thy Fri Apr 12 22:09:25 2019 +0200
@@ -473,7 +473,7 @@
all inputs with a norm that is smaller than that radius and to diverge for all inputs with a
norm that is greater.
\<close>
-definition%important conv_radius :: "(nat \<Rightarrow> 'a :: banach) \<Rightarrow> ereal" where
+definition\<^marker>\<open>tag important\<close> conv_radius :: "(nat \<Rightarrow> 'a :: banach) \<Rightarrow> ereal" where
"conv_radius f = inverse (limsup (\<lambda>n. ereal (root n (norm (f n)))))"
lemma conv_radius_cong_weak [cong]: "(\<And>n. f n = g n) \<Longrightarrow> conv_radius f = conv_radius g"
--- a/src/HOL/Analysis/Tagged_Division.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Tagged_Division.thy Fri Apr 12 22:09:25 2019 +0200
@@ -31,7 +31,7 @@
scaleR_cancel_left scaleR_cancel_right scaleR_add_right scaleR_add_left real_vector_class.scaleR_one
-subsection%unimportant \<open>Sundries\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Sundries\<close>
text\<open>A transitive relation is well-founded if all initial segments are finite.\<close>
@@ -132,10 +132,10 @@
subsection \<open>Bounds on intervals where they exist\<close>
-definition%important interval_upperbound :: "('a::euclidean_space) set \<Rightarrow> 'a"
+definition\<^marker>\<open>tag important\<close> interval_upperbound :: "('a::euclidean_space) set \<Rightarrow> 'a"
where "interval_upperbound s = (\<Sum>i\<in>Basis. (SUP x\<in>s. x\<bullet>i) *\<^sub>R i)"
-definition%important interval_lowerbound :: "('a::euclidean_space) set \<Rightarrow> 'a"
+definition\<^marker>\<open>tag important\<close> interval_lowerbound :: "('a::euclidean_space) set \<Rightarrow> 'a"
where "interval_lowerbound s = (\<Sum>i\<in>Basis. (INF x\<in>s. x\<bullet>i) *\<^sub>R i)"
lemma interval_upperbound[simp]:
@@ -194,7 +194,7 @@
subsection \<open>The notion of a gauge --- simply an open set containing the point\<close>
-definition%important "gauge \<gamma> \<longleftrightarrow> (\<forall>x. x \<in> \<gamma> x \<and> open (\<gamma> x))"
+definition\<^marker>\<open>tag important\<close> "gauge \<gamma> \<longleftrightarrow> (\<forall>x. x \<in> \<gamma> x \<and> open (\<gamma> x))"
lemma gaugeI:
assumes "\<And>x. x \<in> \<gamma> x"
@@ -251,7 +251,7 @@
subsection \<open>Divisions\<close>
-definition%important division_of (infixl "division'_of" 40)
+definition\<^marker>\<open>tag important\<close> division_of (infixl "division'_of" 40)
where
"s division_of i \<longleftrightarrow>
finite s \<and>
@@ -997,7 +997,7 @@
subsection \<open>Tagged (partial) divisions\<close>
-definition%important tagged_partial_division_of (infixr "tagged'_partial'_division'_of" 40)
+definition\<^marker>\<open>tag important\<close> tagged_partial_division_of (infixr "tagged'_partial'_division'_of" 40)
where "s tagged_partial_division_of i \<longleftrightarrow>
finite s \<and>
(\<forall>x K. (x, K) \<in> s \<longrightarrow> x \<in> K \<and> K \<subseteq> i \<and> (\<exists>a b. K = cbox a b)) \<and>
@@ -1014,7 +1014,7 @@
(x2, K2) \<in> s \<Longrightarrow> (x1, K1) \<noteq> (x2, K2) \<Longrightarrow> interior K1 \<inter> interior K2 = {}"
using assms unfolding tagged_partial_division_of_def by blast+
-definition%important tagged_division_of (infixr "tagged'_division'_of" 40)
+definition\<^marker>\<open>tag important\<close> tagged_division_of (infixr "tagged'_division'_of" 40)
where "s tagged_division_of i \<longleftrightarrow> s tagged_partial_division_of i \<and> (\<Union>{K. \<exists>x. (x,K) \<in> s} = i)"
lemma tagged_division_of_finite: "s tagged_division_of i \<Longrightarrow> finite s"
@@ -1292,9 +1292,9 @@
\<open>operative_division\<close>. Instances for the monoid are \<^typ>\<open>'a option\<close>, \<^typ>\<open>real\<close>, and
\<^typ>\<open>bool\<close>.\<close>
-paragraph%important \<open>Using additivity of lifted function to encode definedness.\<close>
+paragraph\<^marker>\<open>tag important\<close> \<open>Using additivity of lifted function to encode definedness.\<close>
text \<open>%whitespace\<close>
-definition%important lift_option :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a option \<Rightarrow> 'b option \<Rightarrow> 'c option"
+definition\<^marker>\<open>tag important\<close> lift_option :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a option \<Rightarrow> 'b option \<Rightarrow> 'c option"
where
"lift_option f a' b' = Option.bind a' (\<lambda>a. Option.bind b' (\<lambda>b. Some (f a b)))"
@@ -1304,7 +1304,7 @@
"lift_option f a' None = None"
by (auto simp: lift_option_def)
-lemma%important comm_monoid_lift_option:
+lemma\<^marker>\<open>tag important\<close> comm_monoid_lift_option:
assumes "comm_monoid f z"
shows "comm_monoid (lift_option f) (Some z)"
proof -
@@ -1335,7 +1335,7 @@
paragraph \<open>Division points\<close>
text \<open>%whitespace\<close>
-definition%important "division_points (k::('a::euclidean_space) set) d =
+definition\<^marker>\<open>tag important\<close> "division_points (k::('a::euclidean_space) set) d =
{(j,x). j \<in> Basis \<and> (interval_lowerbound k)\<bullet>j < x \<and> x < (interval_upperbound k)\<bullet>j \<and>
(\<exists>i\<in>d. (interval_lowerbound i)\<bullet>j = x \<or> (interval_upperbound i)\<bullet>j = x)}"
@@ -1858,7 +1858,7 @@
subsection \<open>Fine-ness of a partition w.r.t. a gauge\<close>
-definition%important fine (infixr "fine" 46)
+definition\<^marker>\<open>tag important\<close> fine (infixr "fine" 46)
where "d fine s \<longleftrightarrow> (\<forall>(x,k) \<in> s. k \<subseteq> d x)"
lemma fineI:
@@ -1907,7 +1907,7 @@
qed
(* FIXME structure here, do not have one lemma in each subsection *)
-subsection%unimportant \<open>The set we're concerned with must be closed\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>The set we're concerned with must be closed\<close>
lemma division_of_closed:
fixes i :: "'n::euclidean_space set"
@@ -2301,7 +2301,7 @@
with ptag that show ?thesis by auto
qed
-subsubsection%important \<open>Covering lemma\<close>
+subsubsection\<^marker>\<open>tag important\<close> \<open>Covering lemma\<close>
text\<open> Some technical lemmas used in the approximation results that follow. Proof of the covering
lemma is an obvious multidimensional generalization of Lemma 3, p65 of Swartz's
@@ -2575,7 +2575,7 @@
text \<open>Divisions over all gauges towards finer divisions.\<close>
-definition%important division_filter :: "'a::euclidean_space set \<Rightarrow> ('a \<times> 'a set) set filter"
+definition\<^marker>\<open>tag important\<close> division_filter :: "'a::euclidean_space set \<Rightarrow> ('a \<times> 'a set) set filter"
where "division_filter s = (INF g\<in>{g. gauge g}. principal {p. p tagged_division_of s \<and> g fine p})"
proposition eventually_division_filter:
--- a/src/HOL/Analysis/Topology_Euclidean_Space.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy Fri Apr 12 22:09:25 2019 +0200
@@ -104,7 +104,7 @@
qed
qed
-subsection%unimportant\<open>Balls in Euclidean Space\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Balls in Euclidean Space\<close>
lemma cball_subset_cball_iff:
fixes a :: "'a :: euclidean_space"
@@ -382,11 +382,11 @@
corollary Zero_neq_One[iff]: "0 \<noteq> One"
by (metis One_non_0)
-definition%important (in euclidean_space) eucl_less (infix "<e" 50)
+definition\<^marker>\<open>tag important\<close> (in euclidean_space) eucl_less (infix "<e" 50)
where "eucl_less a b \<longleftrightarrow> (\<forall>i\<in>Basis. a \<bullet> i < b \<bullet> i)"
-definition%important box_eucl_less: "box a b = {x. a <e x \<and> x <e b}"
-definition%important "cbox a b = {x. \<forall>i\<in>Basis. a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i}"
+definition\<^marker>\<open>tag important\<close> box_eucl_less: "box a b = {x. a <e x \<and> x <e b}"
+definition\<^marker>\<open>tag important\<close> "cbox a b = {x. \<forall>i\<in>Basis. a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i}"
lemma box_def: "box a b = {x. \<forall>i\<in>Basis. a \<bullet> i < x \<bullet> i \<and> x \<bullet> i < b \<bullet> i}"
and in_box_eucl_less: "x \<in> box a b \<longleftrightarrow> a <e x \<and> x <e b"
@@ -943,7 +943,7 @@
subsection \<open>General Intervals\<close>
-definition%important "is_interval (s::('a::euclidean_space) set) \<longleftrightarrow>
+definition\<^marker>\<open>tag important\<close> "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)"
lemma is_interval_1:
@@ -1119,7 +1119,7 @@
by (metis image_cong uminus_add_conv_diff)
-subsection%unimportant \<open>Bounded Projections\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Bounded Projections\<close>
lemma bounded_inner_imp_bdd_above:
assumes "bounded s"
@@ -1132,7 +1132,7 @@
by (simp add: assms bounded_imp_bdd_below bounded_linear_image bounded_linear_inner_left)
-subsection%unimportant \<open>Structural rules for pointwise continuity\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Structural rules for pointwise continuity\<close>
lemma continuous_infnorm[continuous_intros]:
"continuous F f \<Longrightarrow> continuous F (\<lambda>x. infnorm (f x))"
@@ -1145,7 +1145,7 @@
using assms unfolding continuous_def by (rule tendsto_inner)
-subsection%unimportant \<open>Structural rules for setwise continuity\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Structural rules for setwise continuity\<close>
lemma continuous_on_infnorm[continuous_intros]:
"continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. infnorm (f x))"
@@ -1160,7 +1160,7 @@
by (rule bounded_bilinear.continuous_on)
-subsection%unimportant \<open>Openness of halfspaces.\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Openness of halfspaces.\<close>
lemma open_halfspace_lt: "open {x. inner a x < b}"
by (simp add: open_Collect_less continuous_on_inner continuous_on_const continuous_on_id)
@@ -1185,7 +1185,7 @@
shows "open {x. x <e a}" "open {x. a <e x}"
by (auto simp: eucl_less_eq_halfspaces open_halfspace_component_lt open_halfspace_component_gt)
-subsection%unimportant \<open>Closure of halfspaces and hyperplanes\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Closure of halfspaces and hyperplanes\<close>
lemma continuous_at_inner: "continuous (at x) (inner a)"
unfolding continuous_at by (intro tendsto_intros)
@@ -1216,7 +1216,7 @@
by (simp add: Collect_ball_eq closed_INT closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id)
-subsection%unimportant\<open>Some more convenient intermediate-value theorem formulations\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Some more convenient intermediate-value theorem formulations\<close>
lemma connected_ivt_hyperplane:
assumes "connected S" and xy: "x \<in> S" "y \<in> S" and b: "inner a x \<le> b" "b \<le> inner a y"
@@ -1537,8 +1537,8 @@
(auto intro!: assms bounded_linear_inner_left bounded_linear_image
simp: euclidean_representation)
-instance%important euclidean_space \<subseteq> heine_borel
-proof%unimportant
+instance\<^marker>\<open>tag important\<close> euclidean_space \<subseteq> heine_borel
+proof
fix f :: "nat \<Rightarrow> 'a"
assume f: "bounded (range f)"
then obtain l::'a and r where r: "strict_mono r"
@@ -1576,7 +1576,7 @@
by auto
qed
-instance%important euclidean_space \<subseteq> banach ..
+instance\<^marker>\<open>tag important\<close> euclidean_space \<subseteq> banach ..
instance euclidean_space \<subseteq> second_countable_topology
proof
@@ -1678,7 +1678,7 @@
qed
-subsection%unimportant\<open>Componentwise limits and continuity\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Componentwise limits and continuity\<close>
text\<open>But is the premise really necessary? Need to generalise @{thm euclidean_dist_l2}\<close>
lemma Euclidean_dist_upper: "i \<in> Basis \<Longrightarrow> dist (x \<bullet> i) (y \<bullet> i) \<le> dist x y"
@@ -1796,7 +1796,7 @@
by (force simp: bounded_linear_def bounded_linear_axioms_def \<open>linear f'\<close>)
qed
-subsection%unimportant \<open>Continuous Extension\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Continuous Extension\<close>
definition clamp :: "'a::euclidean_space \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" where
"clamp a b x = (if (\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i)
@@ -1985,7 +1985,7 @@
qed
-subsection%unimportant \<open>Diameter\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Diameter\<close>
lemma diameter_cball [simp]:
fixes a :: "'a::euclidean_space"
@@ -2046,7 +2046,7 @@
simp: euclidean_dist_l2[where 'a='a] cbox_def dist_norm)
-subsection%unimportant\<open>Relating linear images to open/closed/interior/closure/connected\<close>
+subsection\<^marker>\<open>tag unimportant\<close>\<open>Relating linear images to open/closed/interior/closure/connected\<close>
proposition open_surjective_linear_image:
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space"
@@ -2155,7 +2155,7 @@
using connected_continuous_image assms linear_continuous_on linear_conv_bounded_linear by blast
-subsection%unimportant \<open>"Isometry" (up to constant bounds) of Injective Linear Map\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>"Isometry" (up to constant bounds) of Injective Linear Map\<close>
proposition injective_imp_isometric:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
@@ -2243,7 +2243,7 @@
qed
-subsection%unimportant \<open>Some properties of a canonical subspace\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Some properties of a canonical subspace\<close>
lemma closed_substandard: "closed {x::'a::euclidean_space. \<forall>i\<in>Basis. P i \<longrightarrow> x\<bullet>i = 0}"
(is "closed ?A")
--- a/src/HOL/Analysis/Uniform_Limit.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Uniform_Limit.thy Fri Apr 12 22:09:25 2019 +0200
@@ -12,10 +12,10 @@
subsection \<open>Definition\<close>
-definition%important uniformly_on :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b::metric_space) \<Rightarrow> ('a \<Rightarrow> 'b) filter"
+definition\<^marker>\<open>tag important\<close> uniformly_on :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b::metric_space) \<Rightarrow> ('a \<Rightarrow> 'b) filter"
where "uniformly_on S l = (INF e\<in>{0 <..}. principal {f. \<forall>x\<in>S. dist (f x) (l x) < e})"
-abbreviation%important
+abbreviation\<^marker>\<open>tag important\<close>
"uniform_limit S f l \<equiv> filterlim f (uniformly_on S l)"
definition uniformly_convergent_on where
@@ -366,7 +366,7 @@
qed
text\<open>Alternative version, formulated as in HOL Light\<close>
-corollary%unimportant series_comparison_uniform:
+corollary\<^marker>\<open>tag unimportant\<close> series_comparison_uniform:
fixes f :: "_ \<Rightarrow> nat \<Rightarrow> _ :: banach"
assumes g: "summable g" and le: "\<And>n x. N \<le> n \<and> x \<in> A \<Longrightarrow> norm(f x n) \<le> g n"
shows "\<exists>l. \<forall>e. 0 < e \<longrightarrow> (\<exists>N. \<forall>n x. N \<le> n \<and> x \<in> A \<longrightarrow> dist(sum (f x) {..<n}) (l x) < e)"
@@ -379,20 +379,20 @@
done
qed
-corollary%unimportant Weierstrass_m_test:
+corollary\<^marker>\<open>tag unimportant\<close> Weierstrass_m_test:
fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _ :: banach"
assumes "\<And>n x. x \<in> A \<Longrightarrow> norm (f n x) \<le> M n"
assumes "summable M"
shows "uniform_limit A (\<lambda>n x. \<Sum>i<n. f i x) (\<lambda>x. suminf (\<lambda>i. f i x)) sequentially"
using assms by (intro Weierstrass_m_test_ev always_eventually) auto
-corollary%unimportant Weierstrass_m_test'_ev:
+corollary\<^marker>\<open>tag unimportant\<close> Weierstrass_m_test'_ev:
fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _ :: banach"
assumes "eventually (\<lambda>n. \<forall>x\<in>A. norm (f n x) \<le> M n) sequentially" "summable M"
shows "uniformly_convergent_on A (\<lambda>n x. \<Sum>i<n. f i x)"
unfolding uniformly_convergent_on_def by (rule exI, rule Weierstrass_m_test_ev[OF assms])
-corollary%unimportant Weierstrass_m_test':
+corollary\<^marker>\<open>tag unimportant\<close> Weierstrass_m_test':
fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _ :: banach"
assumes "\<And>n x. x \<in> A \<Longrightarrow> norm (f n x) \<le> M n" "summable M"
shows "uniformly_convergent_on A (\<lambda>n x. \<Sum>i<n. f i x)"
@@ -402,7 +402,7 @@
by simp
-subsection%unimportant \<open>Structural introduction rules\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>Structural introduction rules\<close>
named_theorems uniform_limit_intros "introduction rules for uniform_limit"
setup \<open>
--- a/src/HOL/Analysis/Weierstrass_Theorems.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Weierstrass_Theorems.thy Fri Apr 12 22:09:25 2019 +0200
@@ -8,7 +8,7 @@
subsection \<open>Bernstein polynomials\<close>
-definition%important Bernstein :: "[nat,nat,real] \<Rightarrow> real" where
+definition\<^marker>\<open>tag important\<close> Bernstein :: "[nat,nat,real] \<Rightarrow> real" where
"Bernstein n k x \<equiv> of_nat (n choose k) * x ^ k * (1 - x) ^ (n - k)"
lemma Bernstein_nonneg: "\<lbrakk>0 \<le> x; x \<le> 1\<rbrakk> \<Longrightarrow> 0 \<le> Bernstein n k x"
@@ -201,7 +201,7 @@
lemma prod: "\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> f i \<in> R\<rbrakk> \<Longrightarrow> (\<lambda>x. \<Prod>i \<in> I. f i x) \<in> R"
by (induct I rule: finite_induct; simp add: const mult)
- definition%important normf :: "('a::t2_space \<Rightarrow> real) \<Rightarrow> real"
+ definition\<^marker>\<open>tag important\<close> normf :: "('a::t2_space \<Rightarrow> real) \<Rightarrow> real"
where "normf f \<equiv> SUP x\<in>S. \<bar>f x\<bar>"
lemma normf_upper: "\<lbrakk>continuous_on S f; x \<in> S\<rbrakk> \<Longrightarrow> \<bar>f x\<bar> \<le> normf f"
@@ -792,7 +792,7 @@
declare real_polynomial_function.intros [intro]
-definition%important polynomial_function :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> bool"
+definition\<^marker>\<open>tag important\<close> polynomial_function :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> bool"
where
"polynomial_function p \<equiv> (\<forall>f. bounded_linear f \<longrightarrow> real_polynomial_function (f o p))"
--- a/src/HOL/Analysis/Winding_Numbers.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Analysis/Winding_Numbers.thy Fri Apr 12 22:09:25 2019 +0200
@@ -839,7 +839,7 @@
by (intro closed_segmentI[of u]) (auto simp: u algebra_simps scaleR_conv_of_real complex_eq_iff)
qed
-definition%important rectpath where
+definition\<^marker>\<open>tag important\<close> rectpath where
"rectpath a1 a3 = (let a2 = Complex (Re a3) (Im a1); a4 = Complex (Re a1) (Im a3)
in linepath a1 a2 +++ linepath a2 a3 +++ linepath a3 a4 +++ linepath a4 a1)"
@@ -929,7 +929,7 @@
path_image_rectpath_subset_cbox) simp_all
text\<open>A per-function version for continuous logs, a kind of monodromy\<close>
-proposition%unimportant winding_number_compose_exp:
+proposition\<^marker>\<open>tag unimportant\<close> winding_number_compose_exp:
assumes "path p"
shows "winding_number (exp \<circ> p) 0 = (pathfinish p - pathstart p) / (2 * of_real pi * \<i>)"
proof -
@@ -1020,7 +1020,7 @@
finally show ?thesis .
qed
-subsection%unimportant \<open>The winding number defines a continuous logarithm for the path itself\<close>
+subsection\<^marker>\<open>tag unimportant\<close> \<open>The winding number defines a continuous logarithm for the path itself\<close>
lemma winding_number_as_continuous_log:
assumes "path p" and \<zeta>: "\<zeta> \<notin> path_image p"
--- a/src/HOL/Homology/Invariance_of_Domain.thy Fri Apr 12 19:48:29 2019 +0200
+++ b/src/HOL/Homology/Invariance_of_Domain.thy Fri Apr 12 22:09:25 2019 +0200
@@ -2324,7 +2324,7 @@
assumes contf: "continuous_on S f" and "open S"
and injf: "inj_on f S" and "S \<noteq> {}"
shows "DIM('a) \<le> DIM('b)"
- using%unimportant invariance_of_dimension_subspaces [of UNIV S UNIV f] assms
+ using invariance_of_dimension_subspaces [of UNIV S UNIV f] assms
by auto
corollary continuous_injective_image_subspace_dim_le:
@@ -2334,7 +2334,7 @@
and injf: "inj_on f S"
shows "dim S \<le> dim T"
apply (rule invariance_of_dimension_subspaces [of S S _ f])
- using%unimportant assms by (auto simp: subspace_affine)
+ using assms by (auto simp: subspace_affine)
lemma invariance_of_dimension_convex_domain:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
@@ -2438,8 +2438,8 @@
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "open S" "continuous_on S f" "DIM('b) \<le> DIM('a)" "inj_on f S"
shows "S homeomorphic (f ` S)"
- using%unimportant invariance_of_domain_homeomorphism [OF assms]
- by%unimportant (meson homeomorphic_def)
+ using invariance_of_domain_homeomorphism [OF assms]
+ by (meson homeomorphic_def)
lemma continuous_image_subset_interior:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"