# HG changeset patch # User wenzelm # Date 1555099765 -7200 # Node ID f03a01a18c6ee3d4cfa815adb8421d4a1c30099a # Parent ad6d4a14adb51ab113969763ad33b47cf3377eda modernized tags: default scope excludes proof; diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Abstract_Topology.thy --- 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 \General notion of a topology as a value\ -definition%important istopology :: "('a set \ bool) \ bool" where +definition\<^marker>\tag important\ istopology :: "('a set \ bool) \ bool" where "istopology L \ L {} \ (\S T. L S \ L T \ L (S \ T)) \ (\K. Ball K L \ L (\K))" -typedef%important 'a topology = "{L::('a set) \ bool. istopology L}" +typedef\<^marker>\tag important\ 'a topology = "{L::('a set) \ bool. istopology L}" morphisms "openin" "topology" unfolding istopology_def by blast @@ -116,7 +116,7 @@ subsubsection \Closed sets\ -definition%important closedin :: "'a topology \ 'a set \ bool" where +definition\<^marker>\tag important\ closedin :: "'a topology \ 'a set \ bool" where "closedin U S \ S \ topspace U \ openin U (topspace U - S)" lemma closedin_subset: "closedin U S \ S \ topspace U" @@ -240,7 +240,7 @@ subsection \Subspace topology\ -definition%important subtopology :: "'a topology \ 'a set \ 'a topology" where +definition\<^marker>\tag important\ subtopology :: "'a topology \ 'a set \ 'a topology" where "subtopology U V = topology (\T. \S. T = S \ V \ openin U S)" lemma istopology_subtopology: "istopology (\T. \S. T = S \ V \ openin U S)" @@ -415,7 +415,7 @@ subsection \The canonical topology from the underlying type class\ -abbreviation%important euclidean :: "'a::topological_space topology" +abbreviation\<^marker>\tag important\ euclidean :: "'a::topological_space topology" where "euclidean \ topology open" abbreviation top_of_set :: "'a::topological_space set \ 'a topology" @@ -1416,7 +1416,7 @@ by (simp add: Sup_le_iff closure_of_minimal) -subsection%important \Continuous maps\ +subsection\<^marker>\tag important\ \Continuous maps\ text \We will need to deal with continuous maps in terms of topologies and not in terms of type classes, as defined below.\ @@ -1728,7 +1728,7 @@ declare continuous_map_id_subt [unfolded id_def, simp] -lemma%important continuous_map_alt: +lemma\<^marker>\tag important\ continuous_map_alt: "continuous_map T1 T2 f = ((\U. openin T2 U \ openin T1 (f -` U \ topspace T1)) \ f ` topspace T1 \ 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 \Half-global and completely global cases\ +subsection\<^marker>\tag unimportant\ \Half-global and completely global cases\ 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 \The topology generated by some (open) subsets\ +subsection\<^marker>\tag important\ \The topology generated by some (open) subsets\ text \In the definition below of a generated topology, the \Empty\ case is not necessary, as it follows from \UN\ taking for \K\ 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) \ ('a topology)" +abbreviation\<^marker>\tag unimportant\ topology_generated_by::"('a set set) \ ('a topology)" where "topology_generated_by S \ 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 \Pullback topology\ +subsection\<^marker>\tag important\ \Pullback topology\ text \Pulling back a topology by map gives again a topology. \subtopology\ is a special case of this notion, pulling back by the identity. We introduce the general notion as @@ -4264,7 +4264,7 @@ text \\pullback_topology A f T\ is the pullback of the topology \T\ by the map \f\ on the set \A\.\ -definition%important pullback_topology::"('a set) \ ('a \ 'b) \ ('b topology) \ ('a topology)" +definition\<^marker>\tag important\ pullback_topology::"('a set) \ ('a \ 'b) \ ('b topology) \ ('a topology)" where "pullback_topology A f T = topology (\S. \U. openin T U \ S = f-`U \ A)" lemma istopology_pullback_topology: diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Abstract_Topology_2.thy --- 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 \x = f y\ show "x \ f ` interior S" .. qed -subsection%unimportant \Equality of continuous functions on closure and related results\ +subsection\<^marker>\tag unimportant\ \Equality of continuous functions on closure and related results\ lemma continuous_closedin_preimage_constant: fixes f :: "_ \ 'b::t1_space" @@ -319,7 +319,7 @@ then show ?thesis by auto qed -subsection%unimportant \A function constant on a set\ +subsection\<^marker>\tag unimportant\ \A function constant on a set\ definition constant_on (infixl "(constant'_on)" 50) where "f constant_on A \ \y. \x\A. f x = y" @@ -341,7 +341,7 @@ by metis -subsection%unimportant \Continuity relative to a union.\ +subsection\<^marker>\tag unimportant\ \Continuity relative to a union.\ lemma continuous_on_Un_local: "\closedin (top_of_set (s \ t)) s; closedin (top_of_set (s \ t)) t; @@ -391,7 +391,7 @@ by (auto simp: continuous_on_id intro: continuous_on_cases_le [where h = id, simplified]) -subsection%unimportant\Inverse function property for open/closed maps\ +subsection\<^marker>\tag unimportant\\Inverse function property for open/closed maps\ 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 \Seperability\ +subsection\<^marker>\tag unimportant\ \Seperability\ lemma subset_second_countable: obtains \ :: "'a:: second_countable_topology set set" @@ -540,7 +540,7 @@ qed -subsection%unimportant\Closed Maps\ +subsection\<^marker>\tag unimportant\\Closed Maps\ lemma continuous_imp_closed_map: fixes f :: "'a::t2_space \ 'b::t2_space" @@ -561,7 +561,7 @@ by (fastforce simp add: closedin_closed) qed -subsection%unimportant\Open Maps\ +subsection\<^marker>\tag unimportant\\Open Maps\ lemma open_map_restrict: assumes opeU: "openin (top_of_set (S \ f -` T')) U" @@ -576,7 +576,7 @@ qed -subsection%unimportant\Quotient maps\ +subsection\<^marker>\tag unimportant\\Quotient maps\ lemma quotient_map_imp_continuous_open: assumes T: "f ` S \ 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\Pasting lemmas for functions, for of casewise definitions\ +subsection\<^marker>\tag unimportant\\Pasting lemmas for functions, for of casewise definitions\ subsubsection\on open sets\ @@ -905,11 +905,11 @@ subsection \Retractions\ -definition%important retraction :: "('a::topological_space) set \ 'a set \ ('a \ 'a) \ bool" +definition\<^marker>\tag important\ retraction :: "('a::topological_space) set \ 'a set \ ('a \ 'a) \ bool" where "retraction S T r \ T \ S \ continuous_on S r \ r ` S \ T \ (\x\T. r x = x)" -definition%important retract_of (infixl "retract'_of" 50) where +definition\<^marker>\tag important\ retract_of (infixl "retract'_of" 50) where "T retract_of S \ (\r. retraction S T r)" lemma retraction_idempotent: "retraction S T r \ x \ S \ r (r x) = r x" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Arcwise_Connected.thy --- 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\Arcwise Connections\(*FIX ME this subsection is empty(?) *) +subsection\<^marker>\tag unimportant\\Arcwise Connections\(*FIX ME this subsection is empty(?) *) subsection\Density of points with dyadic rational coordinates\ @@ -201,7 +201,7 @@ qed -definition%unimportant dyadics :: "'a::field_char_0 set" where "dyadics \ \k m. {of_nat m / 2^k}" +definition\<^marker>\tag unimportant\ dyadics :: "'a::field_char_0 set" where "dyadics \ \k m. {of_nat m / 2^k}" lemma real_in_dyadics [simp]: "real m \ dyadics" apply (simp add: dyadics_def) @@ -344,7 +344,7 @@ qed -function%unimportant (domintros) dyad_rec :: "[nat \ 'a, 'a\'a, 'a\'a, real] \ 'a" where +function\<^marker>\tag unimportant\ (domintros) dyad_rec :: "[nat \ 'a, 'a\'a, 'a\'a, real] \ '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>\tag unimportant\ dyad_rec2 where "dyad_rec2 u v lc rc x = dyad_rec (\z. (u,v)) (\(a,b). (a, lc a b (midpoint a b))) (\(a,b). (rc a b (midpoint a b), b)) (2*x)" -abbreviation%unimportant leftrec where "leftrec u v lc rc x \ fst (dyad_rec2 u v lc rc x)" -abbreviation%unimportant rightrec where "rightrec u v lc rc x \ snd (dyad_rec2 u v lc rc x)" +abbreviation\<^marker>\tag unimportant\ leftrec where "leftrec u v lc rc x \ fst (dyad_rec2 u v lc rc x)" +abbreviation\<^marker>\tag unimportant\ rightrec where "rightrec u v lc rc x \ 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) diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Ball_Volume.thy --- 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. \ -definition%important unit_ball_vol :: "real \ real" where +definition\<^marker>\tag important\ unit_ball_vol :: "real \ real" where "unit_ball_vol n = pi powr (n / 2) / Gamma (n / 2 + 1)" lemma unit_ball_vol_pos [simp]: "n \ 0 \ unit_ball_vol n > 0" @@ -177,7 +177,7 @@ assumes r: "r \ 0" begin -theorem%unimportant emeasure_cball: +theorem\<^marker>\tag unimportant\ 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>\tag unimportant\ 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>\tag unimportant\ 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 \ null_sets lborel" @@ -219,7 +219,7 @@ finally show ?thesis .. qed -corollary%important content_ball: +corollary\<^marker>\tag important\ 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>\tag unimportant\ 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>\tag unimportant\ circle_area: "r \ 0 \ 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>\tag unimportant\ sphere_volume: "r \ 0 \ content (ball c r :: (real ^ 3) set) = 4 / 3 * r ^ 3 * pi" by (simp add: content_ball unit_ball_vol_3) text \ Useful equivalent forms \ -corollary%unimportant content_ball_eq_0_iff [simp]: "content (ball c r) = 0 \ r \ 0" +corollary\<^marker>\tag unimportant\ content_ball_eq_0_iff [simp]: "content (ball c r) = 0 \ r \ 0" proof - have "r > 0 \ 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) \ 0 < r" +corollary\<^marker>\tag unimportant\ content_ball_gt_0_iff [simp]: "0 < content (ball z r) \ 0 < r" by (auto simp: zero_less_measure_iff) -corollary%unimportant content_cball_eq_0_iff [simp]: "content (cball c r) = 0 \ r \ 0" +corollary\<^marker>\tag unimportant\ content_cball_eq_0_iff [simp]: "content (cball c r) = 0 \ r \ 0" proof (cases "r = 0") case False moreover have "r > 0 \ 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) \ 0 < r" +corollary\<^marker>\tag unimportant\ content_cball_gt_0_iff [simp]: "0 < content (cball z r) \ 0 < r" by (auto simp: zero_less_measure_iff) end diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Binary_Product_Measure.thy --- 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 "\\<^sub>M" 80) where +definition\<^marker>\tag important\ pair_measure (infixr "\\<^sub>M" 80) where "A \\<^sub>M B = measure_of (space A \ space B) {a \ b | a b. a \ sets A \ b \ sets B} (\X. \\<^sup>+x. (\\<^sup>+y. indicator X (x,y) \B) \A)" @@ -338,7 +338,7 @@ subsection \Binary products of \\\-finite emeasure spaces\ -locale%unimportant pair_sigma_finite = M1?: sigma_finite_measure M1 + M2?: sigma_finite_measure M2 +locale\<^marker>\tag unimportant\ 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 \ P?: sigma_finite_measure "M1 \\<^sub>M M2" +sublocale\<^marker>\tag unimportant\ pair_sigma_finite \ P?: sigma_finite_measure "M1 \\<^sub>M M2" proof from M1.sigma_finite_countable guess F1 .. moreover from M2.sigma_finite_countable guess F2 .. diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Bochner_Integration.thy --- 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 \ space M. f y \ 0} \ \" by simp qed fact -definition%important simple_bochner_integral :: "'a measure \ ('a \ 'b::real_vector) \ 'b" where +definition\<^marker>\tag important\ simple_bochner_integral :: "'a measure \ ('a \ 'b::real_vector) \ 'b" where "simple_bochner_integral M f = (\y\f`space M. measure M {x\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>\tag important\ lebesgue_integral ("integral\<^sup>L") where "integral\<^sup>L M f = (if \x. has_bochner_integral M f x then THE x. has_bochner_integral M f x else 0)" syntax diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Borel_Space.thy --- 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 \ \r s. r \ A \ s \ A \ r \ s \ f r \ f s" +definition\<^marker>\tag important\ "mono_on f A \ \r s. r \ A \ s \ A \ r \ s \ f r \ f s" lemma mono_onI: "(\r s. r \ A \ s \ A \ r \ s \ f r \ f s) \ mono_on f A" @@ -41,7 +41,7 @@ lemma mono_on_subset: "mono_on f A \ B \ A \ mono_on f B" unfolding mono_on_def by auto -definition%important "strict_mono_on f A \ \r s. r \ A \ s \ A \ r < s \ f r < f s" +definition\<^marker>\tag important\ "strict_mono_on f A \ \r s. r \ A \ s \ A \ r < s \ f r < f s" lemma strict_mono_onI: "(\r s. r \ A \ s \ A \ r < s \ f r < f s) \ strict_mono_on f A" @@ -296,7 +296,7 @@ subsection \Generic Borel spaces\ -definition%important (in topological_space) borel :: "'a measure" where +definition\<^marker>\tag important\ (in topological_space) borel :: "'a measure" where "borel = sigma UNIV {S. open S}" abbreviation "borel_measurable M \ measurable M borel" @@ -1508,7 +1508,7 @@ lemma borel_measurable_arctan [measurable]: "arctan \ borel_measurable borel" by (intro borel_measurable_continuous_on1 continuous_intros) -lemma%important borel_measurable_complex_iff: +lemma\<^marker>\tag important\ borel_measurable_complex_iff: "f \ borel_measurable M \ (\x. Re (f x)) \ borel_measurable M \ (\x. Im (f x)) \ borel_measurable M" apply auto @@ -1701,7 +1701,7 @@ "f \ M \\<^sub>M borel \ (\x. enn2real (f x)) \ M \\<^sub>M borel" unfolding enn2real_def[abs_def] by measurable -definition%important [simp]: "is_borel f M \ f \ borel_measurable M" +definition\<^marker>\tag important\ [simp]: "is_borel f M \ f \ 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] diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Bounded_Continuous_Function.thy --- 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 \Definition\ -definition%important "bcontfun = {f. continuous_on UNIV f \ bounded (range f)}" +definition\<^marker>\tag important\ "bcontfun = {f. continuous_on UNIV f \ bounded (range f)}" typedef (overloaded) ('a, 'b) bcontfun ("(_ \\<^sub>C /_)" [22, 21] 21) = "bcontfun::('a::topological_space \ 'b::metric_space) set" @@ -41,7 +41,7 @@ instantiation bcontfun :: (topological_space, metric_space) metric_space begin -lift_definition%important dist_bcontfun :: "'a \\<^sub>C 'b \ 'a \\<^sub>C 'b \ real" +lift_definition\<^marker>\tag important\ dist_bcontfun :: "'a \\<^sub>C 'b \ 'a \\<^sub>C 'b \ real" is "\f g. (SUP x. dist (f x) (g x))" . definition uniformity_bcontfun :: "('a \\<^sub>C 'b \ 'a \\<^sub>C 'b) filter" @@ -175,8 +175,8 @@ subsection \Complete Space\ -instance%important bcontfun :: (metric_space, complete_space) complete_space -proof%unimportant +instance\<^marker>\tag important\ bcontfun :: (metric_space, complete_space) complete_space +proof fix f :: "nat \ ('a, 'b) bcontfun" assume "Cauchy f" \ \Cauchy equals uniform convergence\ then obtain g where "uniform_limit UNIV f g sequentially" @@ -191,9 +191,9 @@ qed -subsection%unimportant \Supremum norm for a normed vector space\ +subsection\<^marker>\tag unimportant\ \Supremum norm for a normed vector space\ -instantiation%unimportant bcontfun :: (topological_space, real_normed_vector) real_vector +instantiation\<^marker>\tag unimportant\ bcontfun :: (topological_space, real_normed_vector) real_vector begin lemma uminus_cont: "f \ bcontfun \ (\x. - f x) \ bcontfun" for f::"'a \ 'b" @@ -247,7 +247,7 @@ "bounded (range f) \ norm (f x) \ (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>\tag unimportant\ bcontfun :: (topological_space, real_normed_vector) real_normed_vector begin definition norm_bcontfun :: "('a, 'b) bcontfun \ real" @@ -290,7 +290,7 @@ using dist_bound[of f 0 b] assms by (simp add: dist_norm) -subsection%unimportant \(bounded) continuous extenstion\ +subsection\<^marker>\tag unimportant\ \(bounded) continuous extenstion\ lemma integral_clamp: "integral {t0::real..clamp t0 t1 x} f = diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Bounded_Linear_Function.thy --- 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 \Intro rules for \<^term>\bounded_linear\\ +subsection\<^marker>\tag unimportant\ \Intro rules for \<^term>\bounded_linear\\ named_theorems bounded_linear_intros @@ -81,7 +81,7 @@ bounded_linear_inner_right_comp -subsection%unimportant \declaration of derivative/continuous/tendsto introduction rules for bounded linear functions\ +subsection\<^marker>\tag unimportant\ \declaration of derivative/continuous/tendsto introduction rules for bounded linear functions\ attribute_setup bounded_linear = \Scan.succeed (Thm.declaration_attribute (fn thm => @@ -116,7 +116,7 @@ subsection \Type of bounded linear functions\ -typedef%important (overloaded) ('a, 'b) blinfun ("(_ \\<^sub>L /_)" [22, 21] 21) = +typedef\<^marker>\tag important\ (overloaded) ('a, 'b) blinfun ("(_ \\<^sub>L /_)" [22, 21] 21) = "{f::'a::real_normed_vector\'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 \\<^sub>L 'b \ real" is onorm . +lift_definition\<^marker>\tag important\ norm_blinfun :: "'a \\<^sub>L 'b \ real" is onorm . lift_definition minus_blinfun :: "'a \\<^sub>L 'b \ 'a \\<^sub>L 'b \ 'a \\<^sub>L 'b" is "\f g x. f x - g x" @@ -160,14 +160,14 @@ lift_definition uminus_blinfun :: "'a \\<^sub>L 'b \ 'a \\<^sub>L 'b" is "\f x. - f x" by (rule bounded_linear_minus) -lift_definition%important zero_blinfun :: "'a \\<^sub>L 'b" is "\x. 0" +lift_definition\<^marker>\tag important\ zero_blinfun :: "'a \\<^sub>L 'b" is "\x. 0" by (rule bounded_linear_zero) -lift_definition%important plus_blinfun :: "'a \\<^sub>L 'b \ 'a \\<^sub>L 'b \ 'a \\<^sub>L 'b" +lift_definition\<^marker>\tag important\ plus_blinfun :: "'a \\<^sub>L 'b \ 'a \\<^sub>L 'b \ 'a \\<^sub>L 'b" is "\f g x. f x + g x" by (metis bounded_linear_add) -lift_definition%important scaleR_blinfun::"real \ 'a \\<^sub>L 'b \ 'a \\<^sub>L 'b" is "\r f x. r *\<^sub>R f x" +lift_definition\<^marker>\tag important\ scaleR_blinfun::"real \ 'a \\<^sub>L 'b \ 'a \\<^sub>L 'b" is "\r f x. r *\<^sub>R f x" by (metis bounded_linear_compose bounded_linear_scaleR_right) definition sgn_blinfun :: "'a \\<^sub>L 'b \ 'a \\<^sub>L 'b" @@ -367,7 +367,7 @@ by (rule convergentI) qed -subsection%unimportant \On Euclidean Space\ +subsection\<^marker>\tag unimportant\ \On Euclidean Space\ lemma Zfun_sum: assumes "finite s" @@ -588,7 +588,7 @@ qed -subsection%unimportant \concrete bounded linear functions\ +subsection\<^marker>\tag unimportant\ \concrete bounded linear functions\ lemma transfer_bounded_bilinear_bounded_linearI: assumes "g = (\i x. (blinfun_apply (f i) x))" @@ -757,7 +757,7 @@ defined analogously. \ -definition%important strong_operator_topology::"('a::real_normed_vector \\<^sub>L'b::real_normed_vector) topology" +definition\<^marker>\tag important\ strong_operator_topology::"('a::real_normed_vector \\<^sub>L'b::real_normed_vector) topology" where "strong_operator_topology = pullback_topology UNIV blinfun_apply euclidean" lemma strong_operator_topology_topspace: diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Brouwer_Fixpoint.thy --- 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."\ -definition%important AR :: "'a::topological_space set \ bool" where +definition\<^marker>\tag important\ AR :: "'a::topological_space set \ bool" where "AR S \ \U. \S'::('a * real) set. S homeomorphic S' \ closedin (top_of_set U) S' \ S' retract_of U" -definition%important ANR :: "'a::topological_space set \ bool" where +definition\<^marker>\tag important\ ANR :: "'a::topological_space set \ bool" where "ANR S \ \U. \S'::('a * real) set. S homeomorphic S' \ closedin (top_of_set U) S' \ (\T. openin (top_of_set U) T \ S' retract_of T)" -definition%important ENR :: "'a::topological_space set \ bool" where +definition\<^marker>\tag important\ ENR :: "'a::topological_space set \ bool" where "ENR S \ \U. open U \ S retract_of U" text \First, show that we do indeed get the "usual" properties of ARs and ANRs.\ @@ -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>\tag unimportant\ 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>\tag unimportant\ nullhomotopic_into_ANR_extension: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "closed S" and contf: "continuous_on S f" @@ -4288,7 +4288,7 @@ by (force elim: homotopic_with_eq [of _ _ _ g "\x. c"] simp: \\x. x \ S \ g x = f x\) qed -corollary%unimportant nullhomotopic_into_rel_frontier_extension: +corollary\<^marker>\tag unimportant\ nullhomotopic_into_rel_frontier_extension: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "closed S" and contf: "continuous_on S f" @@ -4299,7 +4299,7 @@ (\g. continuous_on UNIV g \ range g \ rel_frontier T \ (\x \ 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>\tag unimportant\ nullhomotopic_into_sphere_extension: fixes f :: "'a::euclidean_space \ 'b :: euclidean_space" assumes "closed S" and contf: "continuous_on S f" and "S \ {}" and fim: "f ` S \ sphere a r" @@ -4321,7 +4321,7 @@ done qed -proposition%unimportant Borsuk_map_essential_bounded_component: +proposition\<^marker>\tag unimportant\ Borsuk_map_essential_bounded_component: fixes a :: "'a :: euclidean_space" assumes "compact S" and "a \ S" shows "bounded (connected_component_set (- S) a) \ @@ -4578,7 +4578,7 @@ qed qed -proposition%unimportant homotopic_neighbourhood_extension: +proposition\<^marker>\tag unimportant\ homotopic_neighbourhood_extension: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and fim: "f ` S \ U" and contg: "continuous_on S g" and gim: "g ` S \ U" @@ -4643,7 +4643,7 @@ qed text\ Homotopy on a union of closed-open sets.\ -proposition%unimportant homotopic_on_clopen_Union: +proposition\<^marker>\tag unimportant\ homotopic_on_clopen_Union: fixes \ :: "'a::euclidean_space set set" assumes "\S. S \ \ \ closedin (top_of_set (\\)) S" and "\S. S \ \ \ openin (top_of_set (\\)) S" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Caratheodory.thy --- 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 \Characterizations of Measures\ -definition%important outer_measure_space where +definition\<^marker>\tag important\ outer_measure_space where "outer_measure_space M f \ positive M f \ increasing M f \ countably_subadditive M f" -subsubsection%important \Lambda Systems\ +subsubsection\<^marker>\tag important\ \Lambda Systems\ -definition%important lambda_system :: "'a set \ 'a set set \ ('a set \ ennreal) \ 'a set set" +definition\<^marker>\tag important\ lambda_system :: "'a set \ 'a set set \ ('a set \ ennreal) \ 'a set set" where "lambda_system \ M f = {l \ M. \x \ M. f (l \ x) + f ((\ - l) \ x) = f x}" @@ -297,7 +297,7 @@ using pos by (simp add: measure_space_def) qed -definition%important outer_measure :: "'a set set \ ('a set \ ennreal) \ 'a set \ ennreal" where +definition\<^marker>\tag important\ outer_measure :: "'a set set \ ('a set \ ennreal) \ 'a set \ ennreal" where "outer_measure M f X = (INF A\{A. range A \ M \ disjoint_family A \ X \ (\i. A i)}. \i. f (A i))" @@ -499,7 +499,7 @@ subsection \Volumes\ -definition%important volume :: "'a set set \ ('a set \ ennreal) \ bool" where +definition\<^marker>\tag important\ volume :: "'a set set \ ('a set \ ennreal) \ bool" where "volume M f \ (f {} = 0) \ (\a\M. 0 \ f a) \ (\C\M. disjoint C \ finite C \ \C \ M \ f (\C) = (\c\C. f c))" @@ -635,7 +635,7 @@ qed qed -subsubsection%important \Caratheodory on semirings\ +subsubsection\<^marker>\tag important\ \Caratheodory on semirings\ theorem (in semiring_of_sets) caratheodory: assumes pos: "positive M \" and ca: "countably_additive M \" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Cartesian_Euclidean_Space.thy --- 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>\tag important\ "jacobian f net = matrix(frechet_derivative f net)" proposition jacobian_works: "(f::(real^'a) \ (real^'b)) differentiable net \ @@ -591,7 +591,7 @@ vector_cart[of "\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\Routine results connecting the types \<^typ>\real^1\ and \<^typ>\real\\ +subsection\<^marker>\tag unimportant\\Routine results connecting the types \<^typ>\real^1\ and \<^typ>\real\\ lemma vec_cbox_1_eq [simp]: shows "vec ` cbox u v = cbox (vec u) (vec v ::real^1)" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Cartesian_Space.thy --- 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 \Type @{typ \'a ^ 'n\} and fields as vector spaces\ (*much of the following +subsection\<^marker>\tag unimportant\ \Type @{typ \'a ^ 'n\} and fields as vector spaces\ (*much of the following is really basic linear algebra, check for overlap? rename subsection? *) definition "cart_basis = {axis i 1 | i. i\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>\tag important\ rank :: "'a::field^'n^'m=>nat" where row_rank_def_gen: "rank A \ 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 \Lemmas for working on \real^1/2/3/4\\ +subsection\<^marker>\tag unimportant\ \Lemmas for working on \real^1/2/3/4\\ 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\The collapse of the general concepts to dimension one\ +subsection\<^marker>\tag unimportant\\The collapse of the general concepts to dimension one\ lemma vector_one: "(x::'a ^1) = (\ i. (x$1))" by (simp add: vec_eq_iff) @@ -762,7 +762,7 @@ by (auto simp add: norm_real dist_norm) -subsection%unimportant\Routine results connecting the types \<^typ>\real^1\ and \<^typ>\real\\ +subsection\<^marker>\tag unimportant\\Routine results connecting the types \<^typ>\real^1\ and \<^typ>\real\\ lemma vector_one_nth [simp]: fixes x :: "'a^1" shows "vec (x $ 1) = x" @@ -795,7 +795,7 @@ done -subsection%unimportant\Explicit vector construction from lists\ +subsection\<^marker>\tag unimportant\\Explicit vector construction from lists\ definition "vector l = (\ i. foldr (\x f n. fun_upd (f (n+1)) n x) l (\n x. 0) 1 i)" @@ -835,7 +835,7 @@ apply (simp add: forall_3) done -subsection%unimportant \lambda skolemization on cartesian products\ +subsection\<^marker>\tag unimportant\ \lambda skolemization on cartesian products\ lemma lambda_skolem: "(\i. \x. P i x) \ (\x::'a ^ 'n. \i. P i (x $ i))" (is "?lhs \ ?rhs") @@ -950,7 +950,7 @@ lemma const_vector_cart:"((\ i. d)::real^'n) = (\i\Basis. d *\<^sub>R i)" by (rule vector_cart) -subsection%unimportant \Explicit formulas for low dimensions\ +subsection\<^marker>\tag unimportant\ \Explicit formulas for low dimensions\ lemma prod_neutral_const: "prod f {(1::nat)..1} = f 1" by simp @@ -964,7 +964,7 @@ subsection \Orthogonality of a matrix\ -definition%important "orthogonal_matrix (Q::'a::semiring_1^'n^'n) \ +definition\<^marker>\tag important\ "orthogonal_matrix (Q::'a::semiring_1^'n^'n) \ transpose Q ** Q = mat 1 \ Q ** transpose Q = mat 1" lemma orthogonal_matrix: "orthogonal_matrix (Q:: real ^'n^'n) \ transpose Q ** Q = mat 1" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Cauchy_Integral_Theorem.thy --- 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 \Homeomorphisms of arc images\ +subsection\<^marker>\tag unimportant\ \Homeomorphisms of arc images\ lemma homeomorphism_arc: fixes g :: "real \ 'a::t2_space" @@ -85,7 +85,7 @@ using arc_simple_path inside_arc_empty by blast -subsection%unimportant \Piecewise differentiable functions\ +subsection\<^marker>\tag unimportant\ \Piecewise differentiable functions\ 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.\ -definition%important C1_differentiable_on :: "(real \ 'a::real_normed_vector) \ real set \ bool" +definition\<^marker>\tag important\ C1_differentiable_on :: "(real \ 'a::real_normed_vector) \ real set \ bool" (infix "C1'_differentiable'_on" 50) where "f C1_differentiable_on S \ @@ -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>\tag important\ piecewise_C1_differentiable_on (infixr "piecewise'_C1'_differentiable'_on" 50) where "f piecewise_C1_differentiable_on i \ continuous_on i f \ @@ -704,7 +704,7 @@ subsection \Valid paths, and their start and finish\ -definition%important valid_path :: "(real \ 'a :: real_normed_vector) \ bool" +definition\<^marker>\tag important\ valid_path :: "(real \ 'a :: real_normed_vector) \ bool" where "valid_path f \ f piecewise_C1_differentiable_on {0..1::real}" definition closed_path :: "(real \ 'a :: real_normed_vector) \ bool" @@ -803,17 +803,17 @@ text\piecewise differentiable function on [0,1]\ -definition%important has_contour_integral :: "(complex \ complex) \ complex \ (real \ complex) \ bool" +definition\<^marker>\tag important\ has_contour_integral :: "(complex \ complex) \ complex \ (real \ complex) \ bool" (infixr "has'_contour'_integral" 50) where "(f has_contour_integral i) g \ ((\x. f(g x) * vector_derivative g (at x within {0..1})) has_integral i) {0..1}" -definition%important contour_integrable_on +definition\<^marker>\tag important\ contour_integrable_on (infixr "contour'_integrable'_on" 50) where "f contour_integrable_on g \ \i. (f has_contour_integral i) g" -definition%important contour_integral +definition\<^marker>\tag important\ contour_integral where "contour_integral g f \ SOME i. (f has_contour_integral i) g \ \ f contour_integrable_on g \ i=0" lemma not_integrable_contour_integral: "\ f contour_integrable_on g \ contour_integral g f = 0" @@ -870,7 +870,7 @@ (\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 \Reversing a path\ +subsection\<^marker>\tag unimportant\ \Reversing a path\ lemma valid_path_imp_reverse: assumes "valid_path g" @@ -948,7 +948,7 @@ qed -subsection%unimportant \Joining two paths together\ +subsection\<^marker>\tag unimportant\ \Joining two paths together\ 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 \Shifting the starting point of a (closed) path\ +subsection\<^marker>\tag unimportant\ \Shifting the starting point of a (closed) path\ lemma shiftpath_alt_def: "shiftpath a f = (\x. if x \ 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 \More about straight-line paths\ +subsection\<^marker>\tag unimportant\ \More about straight-line paths\ 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]: "(\x. x) contour_integrable_on (linepath a b)" by (simp add: continuous_on_id contour_integrable_continuous_linepath) -subsection%unimportant \Arithmetical combining theorems\ +subsection\<^marker>\tag unimportant\ \Arithmetical combining theorems\ lemma has_contour_integral_neg: "(f has_contour_integral i) g \ ((\x. -(f x)) has_contour_integral (-i)) g" @@ -1753,7 +1753,7 @@ \ ((\x. sum (\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 \Operations on path integrals\ +subsection\<^marker>\tag unimportant\ \Operations on path integrals\ lemma contour_integral_const_linepath [simp]: "contour_integral (linepath a b) (\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 \Arithmetic theorems for path integrability\ +subsection\<^marker>\tag unimportant\ \Arithmetic theorems for path integrability\ lemma contour_integrable_neg: "f contour_integrable_on g \ (\x. -(f x)) contour_integrable_on g" @@ -1858,7 +1858,7 @@ by (metis has_contour_integral_sum) -subsection%unimportant \Reversing a path integral\ +subsection\<^marker>\tag unimportant\ \Reversing a path integral\ lemma has_contour_integral_reverse_linepath: "(f has_contour_integral i) (linepath a b) @@ -2088,7 +2088,7 @@ qed -subsection%unimportant \The key quadrisection step\ +subsection\<^marker>\tag unimportant\ \The key quadrisection step\ lemma norm_sum_half: assumes "norm(a + b) \ e" @@ -2181,7 +2181,7 @@ qed qed -subsection%unimportant \Cauchy's theorem for triangles\ +subsection\<^marker>\tag unimportant\ \Cauchy's theorem for triangles\ lemma triangle_points_closer: fixes a::complex @@ -2304,7 +2304,7 @@ done qed -proposition%unimportant Cauchy_theorem_triangle: +proposition\<^marker>\tag unimportant\ 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 \Version needing function holomorphic in interior only\ +subsection\<^marker>\tag unimportant\ \Version needing function holomorphic in interior only\ 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 \Version allowing finite number of exceptional points\ - -proposition%unimportant Cauchy_theorem_triangle_cofinite: +subsection\<^marker>\tag unimportant\ \Version allowing finite number of exceptional points\ + +proposition\<^marker>\tag unimportant\ Cauchy_theorem_triangle_cofinite: assumes "continuous_on (convex hull {a,b,c}) f" and "finite S" and "(\x. x \ interior(convex hull {a,b,c}) - S \ f field_differentiable (at x))" @@ -2740,7 +2740,7 @@ qed qed -subsection%unimportant \Cauchy's theorem for an open starlike set\ +subsection\<^marker>\tag unimportant\ \Cauchy's theorem for an open starlike set\ lemma starlike_convex_subset: assumes S: "a \ S" "closed_segment b c \ S" and subs: "\x. x \ S \ closed_segment a x \ S" @@ -2957,7 +2957,7 @@ by (auto intro!: holomorphic_on_imp_differentiable_at simp: interior_open) qed (use assms in \auto intro: holomorphic_on_imp_continuous_on\) -corollary%unimportant Cauchy_theorem_convex: +corollary\<^marker>\tag unimportant\ Cauchy_theorem_convex: "\continuous_on S f; convex S; finite K; \x. x \ interior S - K \ f field_differentiable at x; valid_path g; path_image g \ S; pathfinish g = pathstart g\ @@ -2973,19 +2973,19 @@ using at_within_interior holomorphic_on_def interior_subset by fastforce text\In particular for a disc\ -corollary%unimportant Cauchy_theorem_disc: +corollary\<^marker>\tag unimportant\ Cauchy_theorem_disc: "\finite K; continuous_on (cball a e) f; \x. x \ ball a e - K \ f field_differentiable at x; valid_path g; path_image g \ cball a e; pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" by (auto intro: Cauchy_theorem_convex) -corollary%unimportant Cauchy_theorem_disc_simple: +corollary\<^marker>\tag unimportant\ Cauchy_theorem_disc_simple: "\f holomorphic_on (ball a e); valid_path g; path_image g \ ball a e; pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" by (simp add: Cauchy_theorem_convex_simple) -subsection%unimportant \Generalize integrability to local primitives\ +subsection\<^marker>\tag unimportant\ \Generalize integrability to local primitives\ lemma contour_integral_local_primitive_lemma: fixes f :: "complex\complex" @@ -3422,7 +3422,7 @@ subsection \Winding Numbers\ -definition%important winding_number_prop :: "[real \ complex, complex, real, real \ complex, complex] \ bool" where +definition\<^marker>\tag important\ winding_number_prop :: "[real \ complex, complex, real, real \ complex, complex] \ bool" where "winding_number_prop \ z e p n \ valid_path p \ z \ path_image p \ pathstart p = pathstart \ \ @@ -3430,7 +3430,7 @@ (\t \ {0..1}. norm(\ t - p t) < e) \ contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" -definition%important winding_number:: "[real \ complex, complex] \ complex" where +definition\<^marker>\tag important\ winding_number:: "[real \ complex, complex] \ complex" where "winding_number \ z \ SOME n. \e > 0. \p. winding_number_prop \ z e p n" @@ -3674,7 +3674,7 @@ done qed -subsubsection%unimportant \Some lemmas about negating a path\ +subsubsection\<^marker>\tag unimportant\ \Some lemmas about negating a path\ lemma valid_path_negatepath: "valid_path \ \ valid_path (uminus \ \)" 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 \Useful sufficient conditions for the winding number to be positive\ +subsubsection\<^marker>\tag unimportant\ \Useful sufficient conditions for the winding number to be positive\ lemma Re_winding_number: "\valid_path \; z \ path_image \\ @@ -4228,7 +4228,7 @@ "path \ \ continuous_on (- path_image \) (\w. winding_number \ w)" by (simp add: continuous_at_imp_continuous_on continuous_at_winding_number) -subsection%unimportant \The winding number is constant on a connected region\ +subsection\<^marker>\tag unimportant\ \The winding number is constant on a connected region\ lemma winding_number_constant: assumes \: "path \" and loop: "pathfinish \ = pathstart \" and cs: "connected S" and sg: "S \ path_image \ = {}" @@ -4329,11 +4329,11 @@ finally show ?thesis . qed -corollary%unimportant winding_number_zero_const: "a \ z \ winding_number (\t. a) z = 0" +corollary\<^marker>\tag unimportant\ winding_number_zero_const: "a \ z \ winding_number (\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>\tag unimportant\ winding_number_zero_outside: "\path \; convex s; pathfinish \ = pathstart \; z \ s; path_image \ \ s\ \ winding_number \ 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 \More winding number properties\ +subsection\<^marker>\tag unimportant\ \More winding number properties\ text\including the fact that it's +-1 inside a simple closed curve.\ @@ -4908,7 +4908,7 @@ subsection\Partial circle path\ -definition%important part_circlepath :: "[complex, real, real, real, real] \ complex" +definition\<^marker>\tag important\ part_circlepath :: "[complex, real, real, real, real] \ complex" where "part_circlepath z r s t \ \x. z + of_real r * exp (\ * of_real (linepath s t x))" lemma pathstart_part_circlepath [simp]: @@ -5294,7 +5294,7 @@ subsection\Special case of one complete circle\ -definition%important circlepath :: "[complex, real, real] \ complex" +definition\<^marker>\tag important\ circlepath :: "[complex, real, real] \ complex" where "circlepath z r \ part_circlepath z r 0 (2*pi)" lemma circlepath: "circlepath z r = (\x. z + r * exp(2 * of_real pi * \ * of_real x))" @@ -5487,7 +5487,7 @@ by (simp add: winding_number_circlepath assms) qed -corollary%unimportant Cauchy_integral_circlepath_simple: +corollary\<^marker>\tag unimportant\ Cauchy_integral_circlepath_simple: assumes "f holomorphic_on cball z r" "norm(w - z) < r" shows "((\u. f u/(u - w)) has_contour_integral (2 * of_real pi * \ * f w)) (circlepath z r)" @@ -5584,7 +5584,7 @@ by (rule tendstoI) qed -corollary%unimportant contour_integral_uniform_limit_circlepath: +corollary\<^marker>\tag unimportant\ contour_integral_uniform_limit_circlepath: assumes "\\<^sub>F n::'a in F. (f n) contour_integrable_on (circlepath z r)" and "uniform_limit (sphere z r) f l F" and "\ trivial_limit F" "0 < r" @@ -5593,7 +5593,7 @@ using assms by (auto simp: vector_derivative_circlepath norm_mult intro!: contour_integral_uniform_limit) -subsection%unimportant \General stepping result for derivative formulas\ +subsection\<^marker>\tag unimportant\ \General stepping result for derivative formulas\ lemma Cauchy_next_derivative: assumes "continuous_on (path_image \) f'" @@ -6696,7 +6696,7 @@ qed -subsection%unimportant \Some more simple/convenient versions for applications\ +subsection\<^marker>\tag unimportant\ \Some more simple/convenient versions for applications\ 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>\tag unimportant\ power_series_and_derivative: fixes a :: "nat \ complex" and r::real assumes "summable (\n. a n * r^n)" obtains g g' where "\z \ 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>\tag unimportant\ power_series_holomorphic: assumes "\w. w \ ball z r \ ((\n. a n*(w - z)^n) sums f w)" shows "f holomorphic_on ball z r" proof - @@ -6965,7 +6965,7 @@ (\w \ ball z r. (\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 \Equality between holomorphic functions, on open ball then connected set\ +subsection\<^marker>\tag unimportant\ \Equality between holomorphic functions, on open ball then connected set\ lemma holomorphic_fun_eq_on_ball: "\f holomorphic_on ball z r; g holomorphic_on ball z r; @@ -7048,7 +7048,7 @@ by (subst higher_deriv_diff) (use assms in \auto intro: holomorphic_intros\) qed (use assms in auto) -subsection%unimportant \Some basic lemmas about poles/singularities\ +subsection\<^marker>\tag unimportant\ \Some basic lemmas about poles/singularities\ lemma pole_lemma: assumes holf: "f holomorphic_on S" and a: "a \ 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>\tag unimportant\ holomorphic_logarithm_exists: assumes A: "convex A" "open A" and f: "f holomorphic_on A" "\x. x \ A \ f x \ 0" and z0: "z0 \ A" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Change_Of_Vars.thy --- 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\\F_sigma\ and \G_delta\ sets.\(*FIX ME mv *) (*https://en.wikipedia.org/wiki/F\_set*) -inductive%important fsigma :: "'a::topological_space set \ bool" where +inductive\<^marker>\tag important\ fsigma :: "'a::topological_space set \ bool" where "(\n::nat. closed (F n)) \ fsigma (\(F ` UNIV))" -inductive%important gdelta :: "'a::topological_space set \ bool" where +inductive\<^marker>\tag important\ gdelta :: "'a::topological_space set \ bool" where "(\n::nat. open (F n)) \ gdelta (\(F ` UNIV))" lemma fsigma_Union_compact: @@ -2100,7 +2100,7 @@ assumes S: "S \ sets lebesgue" and f: "\x. x \ S \ (f has_derivative f' x) (at x within S)" shows "(\x. det(matrix(f' x))) \ 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\The localisation wrt S uses the same argument for many similar results.\ (*See HOL Light's MEASURABLE_ON_LEBESGUE_MEASURABLE_PREIMAGE_BOREL, etc.*) @@ -2949,7 +2949,7 @@ and intS: "(\x. \det (matrix (g' x))\ *\<^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} \ real" and g :: "real^'n::_ \ real^'n::_" @@ -2957,7 +2957,7 @@ and "\x. x \ S \ (g has_derivative g' x) (at x within S)" and "(\x. \det (matrix (g' x))\ * f(g x)) integrable_on S" shows "integral (g ` S) f \ integral S (\x. \det (matrix (g' x))\ * 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\Change-of-variables theorem\ @@ -3576,7 +3576,7 @@ and "inj_on g S" shows "f absolutely_integrable_on (g ` S) \ (\x. \det (matrix (g' x))\ *\<^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} \ real^'n" and g :: "real^'m::_ \ real^'m::_" @@ -3586,8 +3586,8 @@ and disj: "(f absolutely_integrable_on (g ` S) \ (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S)" shows "integral (g ` S) f = integral S (\x. \det (matrix (g' x))\ *\<^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 \ real^'n::{finite,wellorder}" and g :: "real \ real" @@ -3635,7 +3635,7 @@ and inj: "inj_on g S" shows "(f absolutely_integrable_on g ` S \ (\x. \g' x\ *\<^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\Change of variables for integrals: special case of linear function\ diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Complete_Measure.thy --- 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>\tag important\ complete_measure = fixes M :: "'a measure" assumes complete: "\A B. B \ A \ A \ null_sets M \ B \ sets M" -definition%important +definition\<^marker>\tag important\ "split_completion M A p = (if A \ sets M then p = (A, {}) else \N'. A = fst p \ snd p \ fst p \ snd p = {} \ fst p \ sets M \ snd p \ N' \ N' \ null_sets M)" -definition%important +definition\<^marker>\tag important\ "main_part M A = fst (Eps (split_completion M A))" -definition%important +definition\<^marker>\tag important\ "null_part M A = snd (Eps (split_completion M A))" -definition%important completion :: "'a measure \ 'a measure" where +definition\<^marker>\tag important\ completion :: "'a measure \ 'a measure" where "completion M = measure_of (space M) { S \ N |S N N'. S \ sets M \ N' \ null_sets M \ N \ N' } (emeasure M \ main_part M)" @@ -67,9 +67,9 @@ by (intro completionI[of _ "\(S ` UNIV)" "\(N ` UNIV)" "\(N' ` UNIV)"]) auto qed -lemma%important sets_completion: +lemma\<^marker>\tag important\ sets_completion: "sets (completion M) = { S \ N |S N N'. S \ sets M \ N' \ null_sets M \ N \ 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 \ sets M \ A \ sets (completion M)" unfolding sets_completion by force -lemma%important measurable_completion: "f \ M \\<^sub>M N \ f \ completion M \\<^sub>M N" - by%unimportant (auto simp: measurable_def) +lemma\<^marker>\tag important\ measurable_completion: "f \ M \\<^sub>M N \ f \ completion M \\<^sub>M N" + by (auto simp: measurable_def) lemma null_sets_completion: assumes "N' \ null_sets M" "N \ N'" shows "N \ sets (completion M)" using assms by (intro sets_completionI[of N "{}" N N']) auto -lemma%important split_completion: +lemma\<^marker>\tag important\ split_completion: assumes "A \ sets (completion M)" shows "split_completion M A (main_part M A, null_part M A)" -proof%unimportant cases +proof cases assume "A \ 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>\tag important\ emeasure_completion[simp]: assumes S: "S \ 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 ?\ = "emeasure M \ main_part M" show "S \ sets (completion M)" "?\ S = emeasure M (main_part M S) " using S by simp_all show "positive (sets (completion M)) ?\" @@ -283,11 +283,11 @@ qed qed -lemma%important completion_ex_borel_measurable: +lemma\<^marker>\tag important\ completion_ex_borel_measurable: fixes g :: "'a \ ennreal" assumes g: "g \ borel_measurable (completion M)" shows "\g'\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 "\i. \f'. simple_function M f' \ (AE x in M. f i x = f' x)" .. @@ -442,19 +442,19 @@ assumes f: "f \ M \\<^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>\tag important\ semifinite_measure = fixes M :: "'a measure" assumes semifinite: "\A. A \ sets M \ emeasure M A = \ \ \B\sets M. B \ A \ emeasure M B < \" -locale%important locally_determined_measure = semifinite_measure + +locale\<^marker>\tag important\ locally_determined_measure = semifinite_measure + assumes locally_determined: "\A. A \ space M \ (\B. B \ sets M \ emeasure M B < \ \ A \ B \ sets M) \ A \ sets M" -locale%important cld_measure = +locale\<^marker>\tag important\ cld_measure = complete_measure M + locally_determined_measure M for M :: "'a measure" -definition%important outer_measure_of :: "'a measure \ 'a set \ ennreal" +definition\<^marker>\tag important\ outer_measure_of :: "'a measure \ 'a set \ ennreal" where "outer_measure_of M A = (INF B \ {B\sets M. A \ B}. emeasure M B)" lemma outer_measure_of_eq[simp]: "A \ sets M \ 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 \ 'a set \ 'a set \ bool" +definition\<^marker>\tag important\ measurable_envelope :: "'a measure \ 'a set \ 'a set \ bool" where "measurable_envelope M A E \ (A \ E \ E \ sets M \ (\F\sets M. emeasure M (F \ E) = outer_measure_of M (F \ A)))" @@ -618,7 +618,7 @@ by (auto simp: Int_absorb1) qed -lemma%important measurable_envelope_eq2: +lemma\<^marker>\tag important\ measurable_envelope_eq2: assumes "A \ E" "E \ sets M" "emeasure M E < \" shows "measurable_envelope M A E \ (emeasure M E = outer_measure_of M A)" proof safe diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Complex_Analysis_Basics.thy --- 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\General lemmas\ +subsection\<^marker>\tag unimportant\\General lemmas\ lemma nonneg_Reals_cmod_eq_Re: "z \ \\<^sub>\\<^sub>0 \ norm z = Re z" by (simp add: complex_nonneg_Reals_iff cmod_eq_Re) @@ -236,11 +236,11 @@ subsection\Holomorphic functions\ -definition%important holomorphic_on :: "[complex \ complex, complex set] \ bool" +definition\<^marker>\tag important\ holomorphic_on :: "[complex \ complex, complex set] \ bool" (infixl "(holomorphic'_on)" 50) where "f holomorphic_on s \ \x\s. f field_differentiable (at x within s)" -named_theorems%important holomorphic_intros "structural introduction rules for holomorphic_on" +named_theorems\<^marker>\tag important\ holomorphic_intros "structural introduction rules for holomorphic_on" lemma holomorphic_onI [intro?]: "(\x. x \ s \ f field_differentiable (at x within s)) \ f holomorphic_on s" by (simp add: holomorphic_on_def) @@ -498,7 +498,7 @@ by (rule nonzero_deriv_nonconstant [of f "deriv f \" \ S]) (use assms in \auto simp: holomorphic_derivI\) -subsection%unimportant\Caratheodory characterization\ +subsection\<^marker>\tag unimportant\\Caratheodory characterization\ lemma field_differentiable_caratheodory_at: "f field_differentiable (at z) \ @@ -514,10 +514,10 @@ subsection\Analyticity on a set\ -definition%important analytic_on (infixl "(analytic'_on)" 50) +definition\<^marker>\tag important\ analytic_on (infixl "(analytic'_on)" 50) where "f analytic_on S \ \x \ S. \e. 0 < e \ f holomorphic_on (ball x e)" -named_theorems%important analytic_intros "introduction rules for proving analyticity" +named_theorems\<^marker>\tag important\ analytic_intros "introduction rules for proving analyticity" lemma analytic_imp_holomorphic: "f analytic_on S \ 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\Analyticity at a point\ +subsection\<^marker>\tag unimportant\\Analyticity at a point\ lemma analytic_at_ball: "f analytic_on {z} \ (\e. 0 f holomorphic_on ball z e)" @@ -769,7 +769,7 @@ by (force simp add: analytic_at) qed -subsection%unimportant\Combining theorems for derivative with ``analytic at'' hypotheses\ +subsection\<^marker>\tag unimportant\\Combining theorems for derivative with ``analytic at'' hypotheses\ lemma assumes "f analytic_on {z}" "g analytic_on {z}" @@ -805,7 +805,7 @@ "f analytic_on {z} \ deriv (\w. f w * c) z = deriv f z * c" by (auto simp: complex_derivative_mult_at deriv_const analytic_on_const) -subsection%unimportant\Complex differentiation of sequences and series\ +subsection\<^marker>\tag unimportant\\Complex differentiation of sequences and series\ (* 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\Bound theorem\ +subsection\<^marker>\tag unimportant\\Bound theorem\ 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\Inverse function theorem for complex derivatives\ +subsection\<^marker>\tag unimportant\\Inverse function theorem for complex derivatives\ lemma has_field_derivative_inverse_basic: shows "DERIV f (g y) :> f' \ @@ -965,7 +965,7 @@ apply (rule has_derivative_inverse_strong_x [of S g y f]) by auto -subsection%unimportant \Taylor on Complex Numbers\ +subsection\<^marker>\tag unimportant\ \Taylor on Complex Numbers\ lemma sum_Suc_reindex: fixes f :: "nat \ 'a::ab_group_add" @@ -1151,7 +1151,7 @@ qed -subsection%unimportant \Polynomal function extremal theorem, from HOL Light\ +subsection\<^marker>\tag unimportant\ \Polynomal function extremal theorem, from HOL Light\ lemma polyfun_extremal_lemma: (*COMPLEX_POLYFUN_EXTREMAL_LEMMA in HOL Light*) fixes c :: "nat \ 'a::real_normed_div_algebra" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Complex_Transcendental.thy --- 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\Möbius transformations\ (* TODO: Figure out what to do with Möbius transformations *) -definition%important "moebius a b c d = (\z. (a*z+b) / (c*z+d :: 'a :: field))" +definition\<^marker>\tag important\ "moebius a b c d = (\z. (a*z+b) / (c*z+d :: 'a :: field))" theorem moebius_inverse: assumes "a * d \ b * c" "c * z + d \ 0" @@ -64,7 +64,7 @@ by (simp add: norm_power Im_power2) qed -subsection%unimportant\The Exponential Function\ +subsection\<^marker>\tag unimportant\\The Exponential Function\ lemma norm_exp_i_times [simp]: "norm (exp(\ * of_real y)) = 1" by simp @@ -116,7 +116,7 @@ using sums_unique2 by blast qed -corollary%unimportant exp_minus_Euler: "exp(-(\ * z)) = cos(z) - \ * sin(z)" +corollary\<^marker>\tag unimportant\ exp_minus_Euler: "exp(-(\ * z)) = cos(z) - \ * sin(z)" using exp_Euler [of "-z"] by simp @@ -154,7 +154,7 @@ lemma Im_sin_nonneg2: "Re z = pi \ Im z \ 0 \ 0 \ Im (sin z)" by (simp add: Re_sin Im_sin algebra_simps) -subsection%unimportant\Relationships between real and complex trigonometric and hyperbolic functions\ +subsection\<^marker>\tag unimportant\\Relationships between real and complex trigonometric and hyperbolic functions\ lemma real_sin_eq [simp]: "Re(sin(of_real x)) = sin x" by (simp add: sin_of_real) @@ -213,7 +213,7 @@ shows "(\x. cos (f x)) holomorphic_on A" using holomorphic_on_compose[OF assms holomorphic_on_cos] by (simp add: o_def) -subsection%unimportant\More on the Polar Representation of Complex Numbers\ +subsection\<^marker>\tag unimportant\\More on the Polar Representation of Complex Numbers\ 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\Taylor series for complex exponential, sine and cosine\ +subsection\<^marker>\tag unimportant\\Taylor series for complex exponential, sine and cosine\ declare power_Suc [simp del] @@ -859,10 +859,10 @@ subsection\The argument of a complex number (HOL Light version)\ -definition%important is_Arg :: "[complex,real] \ bool" +definition\<^marker>\tag important\ is_Arg :: "[complex,real] \ bool" where "is_Arg z r \ z = of_real(norm z) * exp(\ * of_real r)" -definition%important Arg2pi :: "complex \ real" +definition\<^marker>\tag important\ Arg2pi :: "complex \ real" where "Arg2pi z \ if z = 0 then 0 else THE t. 0 \ t \ t < 2*pi \ is_Arg z t" lemma is_Arg_2pi_iff: "is_Arg z (r + of_int k * (2 * pi)) \ is_Arg z r" @@ -940,7 +940,7 @@ done qed -corollary%unimportant +corollary\<^marker>\tag unimportant\ shows Arg2pi_ge_0: "0 \ Arg2pi z" and Arg2pi_lt_2pi: "Arg2pi z < 2*pi" and Arg2pi_eq: "z = of_real(norm z) * exp(\ * of_real(Arg2pi z))" @@ -1023,7 +1023,7 @@ by blast qed auto -corollary%unimportant Arg2pi_gt_0: +corollary\<^marker>\tag unimportant\ Arg2pi_gt_0: assumes "z \ \\<^sub>\\<^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\Analytic properties of tangent function\ +subsection\<^marker>\tag unimportant\\Analytic properties of tangent function\ 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 \ complex" +definition\<^marker>\tag important\ ln_complex :: "complex \ complex" where "ln_complex \ \z. THE w. exp w = z & -pi < Im(w) & Im(w) \ pi" text\NOTE: within this scope, the constant Ln is not yet available!\ @@ -1189,7 +1189,7 @@ apply auto done -subsection%unimportant\Relation to Real Logarithm\ +subsection\<^marker>\tag unimportant\\Relation to Real Logarithm\ lemma Ln_of_real: assumes "0 < z" @@ -1203,10 +1203,10 @@ using assms by simp qed -corollary%unimportant Ln_in_Reals [simp]: "z \ \ \ Re z > 0 \ ln z \ \" +corollary\<^marker>\tag unimportant\ Ln_in_Reals [simp]: "z \ \ \ Re z > 0 \ ln z \ \" by (auto simp: Ln_of_real elim: Reals_cases) -corollary%unimportant Im_Ln_of_real [simp]: "r > 0 \ Im (ln (of_real r)) = 0" +corollary\<^marker>\tag unimportant\ Im_Ln_of_real [simp]: "r > 0 \ Im (ln (of_real r)) = 0" by (simp add: Ln_of_real) lemma cmod_Ln_Reals [simp]: "z \ \ \ 0 < Re z \ cmod (ln z) = norm (ln (Re z))" @@ -1244,13 +1244,13 @@ lemma Re_Ln [simp]: "z \ 0 \ Re(Ln z) = ln(norm z)" by (metis exp_Ln ln_exp norm_exp_eq_Re) -corollary%unimportant ln_cmod_le: +corollary\<^marker>\tag unimportant\ ln_cmod_le: assumes z: "z \ 0" shows "ln (cmod z) \ 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>\tag unimportant\ exists_complex_root: fixes z :: complex assumes "n \ 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>\tag unimportant\ exists_complex_root_nonzero: fixes z::complex assumes "z \ 0" "n \ 0" obtains w where "w \ 0" "z = w ^ n" by (metis exists_complex_root [of n z] assms power_0_left) -subsection%unimportant\Derivative of Ln away from the branch cut\ +subsection\<^marker>\tag unimportant\\Derivative of Ln away from the branch cut\ lemma assumes "z \ \\<^sub>\\<^sub>0" @@ -1456,7 +1456,7 @@ qed -subsection%unimportant\Quadrant-type results for Ln\ +subsection\<^marker>\tag unimportant\\Quadrant-type results for Ln\ lemma cos_lt_zero_pi: "pi/2 < x \ x < 3*pi/2 \ 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\More Properties of Ln\ +subsection\<^marker>\tag unimportant\\More Properties of Ln\ lemma cnj_Ln: assumes "z \ \\<^sub>\\<^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>\tag unimportant\ Ln_times_simple: "\w \ 0; z \ 0; -pi < Im(Ln w) + Im(Ln z); Im(Ln w) + Im(Ln z) \ pi\ \ Ln(w * z) = Ln(w) + Ln(z)" by (simp add: Ln_times) -corollary%unimportant Ln_times_of_real: +corollary\<^marker>\tag unimportant\ Ln_times_of_real: "\r > 0; z \ 0\ \ 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>\tag unimportant\ Ln_times_Reals: "\r \ Reals; Re r > 0; z \ 0\ \ 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>\tag unimportant\ Ln_divide_of_real: "\r > 0; z \ 0\ \ 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>\tag unimportant\ Ln_prod: fixes f :: "'a \ complex" assumes "finite A" "\x. x \ A \ f x \ 0" shows "\n. Ln (prod f A) = (\x \ A. Ln (f x) + (of_int (n x) * (2 * pi)) * \)" @@ -1753,7 +1753,7 @@ text\Finally: it's is defined for the same interval as the complex logarithm: \(-\,\]\.\ -definition%important Arg :: "complex \ real" where "Arg z \ (if z = 0 then 0 else Im (Ln z))" +definition\<^marker>\tag important\ Arg :: "complex \ real" where "Arg z \ (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 \ \\<^sub>\\<^sub>0" shows "Arg z \ 0" +corollary\<^marker>\tag unimportant\ Arg_ne_0: assumes "z \ \\<^sub>\\<^sub>0" shows "Arg z \ 0" using assms by (auto simp: nonneg_Reals_def Arg_eq_0) lemma Arg_eq_pi_iff: "Arg z = pi \ z \ \ \ 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 \ int" where +definition\<^marker>\tag important\ unwinding :: "complex \ int" where "unwinding z \ THE k. of_int k = (z - Ln(exp z)) / (of_real(2*pi) * \)" lemma unwinding: "of_int (unwinding z) = (z - Ln(exp z)) / (of_real(2*pi) * \)" @@ -2008,7 +2008,7 @@ using unwinding_2pi by (simp add: exp_add) -subsection%unimportant\Relation between Ln and Arg2pi, and hence continuity of Arg2pi\ +subsection\<^marker>\tag unimportant\\Relation between Ln and Arg2pi, and hence continuity of Arg2pi\ 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\Complex Powers\ +subsection\<^marker>\tag unimportant\\Complex Powers\ lemma powr_to_1 [simp]: "z powr 1 = (z::complex)" by (simp add: powr_def) @@ -2535,7 +2535,7 @@ qed -subsection%unimportant\Some Limits involving Logarithms\ +subsection\<^marker>\tag unimportant\\Some Limits involving Logarithms\ lemma lim_Ln_over_power: fixes s::complex @@ -2686,7 +2686,7 @@ qed -subsection%unimportant\Relation between Square Root and exp/ln, hence its derivative\ +subsection\<^marker>\tag unimportant\\Relation between Square Root and exp/ln, hence its derivative\ lemma csqrt_exp_Ln: assumes "z \ 0" @@ -2759,7 +2759,7 @@ "z \ \\<^sub>\\<^sub>0 \ continuous (at z) csqrt" by (simp add: field_differentiable_within_csqrt field_differentiable_imp_continuous_at) -corollary%unimportant isCont_csqrt' [simp]: +corollary\<^marker>\tag unimportant\ isCont_csqrt' [simp]: "\isCont f z; f z \ \\<^sub>\\<^sub>0\ \ isCont (\x. csqrt (f x)) z" by (blast intro: isCont_o2 [OF _ continuous_at_csqrt]) @@ -2802,7 +2802,7 @@ text\The branch cut gives standard bounds in the real case.\ -definition%important Arctan :: "complex \ complex" where +definition\<^marker>\tag important\ Arctan :: "complex \ complex" where "Arctan \ \z. (\/2) * Ln((1 - \*z) / (1 + \*z))" lemma Arctan_def_moebius: "Arctan z = \/2 * Ln (moebius (-\) 1 \ 1 z)" @@ -3051,7 +3051,7 @@ finally show ?thesis by (subst (asm) sums_of_real_iff) qed -subsection%unimportant \Real arctangent\ +subsection\<^marker>\tag unimportant\ \Real arctangent\ 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 \Bounds on pi using real arctangent\ +subsection\<^marker>\tag unimportant\ \Bounds on pi using real arctangent\ lemma pi_machin: "pi = 16 * arctan (1 / 5) - 4 * arctan (1 / 239)" using machin by simp @@ -3229,7 +3229,7 @@ subsection\Inverse Sine\ -definition%important Arcsin :: "complex \ complex" where +definition\<^marker>\tag important\ Arcsin :: "complex \ complex" where "Arcsin \ \z. -\ * Ln(\ * z + csqrt(1 - z\<^sup>2))" lemma Arcsin_body_lemma: "\ * z + csqrt(1 - z\<^sup>2) \ 0" @@ -3396,7 +3396,7 @@ subsection\Inverse Cosine\ -definition%important Arccos :: "complex \ complex" where +definition\<^marker>\tag important\ Arccos :: "complex \ complex" where "Arccos \ \z. -\ * Ln(z + \ * csqrt(1 - z\<^sup>2))" lemma Arccos_range_lemma: "\Re z\ < 1 \ 0 < Im(z + \ * csqrt(1 - z\<^sup>2))" @@ -3555,7 +3555,7 @@ by (simp add: field_differentiable_within_Arccos holomorphic_on_def) -subsection%unimportant\Upper and Lower Bounds for Inverse Sine and Cosine\ +subsection\<^marker>\tag unimportant\\Upper and Lower Bounds for Inverse Sine and Cosine\ lemma Arcsin_bounds: "\Re z\ < 1 \ \Re(Arcsin z)\ < pi/2" unfolding Re_Arcsin @@ -3608,7 +3608,7 @@ qed -subsection%unimportant\Interrelations between Arcsin and Arccos\ +subsection\<^marker>\tag unimportant\\Interrelations between Arcsin and Arccos\ lemma cos_Arcsin_nonzero: assumes "z\<^sup>2 \ 1" shows "cos(Arcsin z) \ 0" @@ -3710,7 +3710,7 @@ by (simp add: Arcsin_Arccos_csqrt_pos) -subsection%unimportant\Relationship with Arcsin on the Real Numbers\ +subsection\<^marker>\tag unimportant\\Relationship with Arcsin on the Real Numbers\ lemma Im_Arcsin_of_real: assumes "\x\ \ 1" @@ -3727,7 +3727,7 @@ by (simp add: Im_Arcsin exp_minus) qed -corollary%unimportant Arcsin_in_Reals [simp]: "z \ \ \ \Re z\ \ 1 \ Arcsin z \ \" +corollary\<^marker>\tag unimportant\ Arcsin_in_Reals [simp]: "z \ \ \ \Re z\ \ 1 \ Arcsin z \ \" 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\Relationship with Arccos on the Real Numbers\ +subsection\<^marker>\tag unimportant\\Relationship with Arccos on the Real Numbers\ lemma Im_Arccos_of_real: assumes "\x\ \ 1" @@ -3777,7 +3777,7 @@ by (simp add: Im_Arccos exp_minus) qed -corollary%unimportant Arccos_in_Reals [simp]: "z \ \ \ \Re z\ \ 1 \ Arccos z \ \" +corollary\<^marker>\tag unimportant\ Arccos_in_Reals [simp]: "z \ \ \ \Re z\ \ 1 \ Arccos z \ \" 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: "\x\ \ 1 \ 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\Some interrelationships among the real inverse trig functions\ +subsection\<^marker>\tag unimportant\\Some interrelationships among the real inverse trig functions\ 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\Continuity results for arcsin and arccos\ +subsection\<^marker>\tag unimportant\\Continuity results for arcsin and arccos\ lemma continuous_on_Arcsin_real [continuous_intros]: "continuous_on {w \ \. \Re w\ \ 1} Arcsin" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Conformal_Mappings.thy --- 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\No need for \<^term>\S\ to be connected. But the nonconstant condition is stronger.\ -corollary%unimportant open_mapping_thm2: +corollary\<^marker>\tag unimportant\ open_mapping_thm2: assumes holf: "f holomorphic_on S" and S: "open S" and "open U" "U \ S" @@ -494,7 +494,7 @@ by force qed -corollary%unimportant open_mapping_thm3: +corollary\<^marker>\tag unimportant\ 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 \\ \ S\ closure_subset by fastforce qed -corollary%unimportant maximum_real_frontier: +corollary\<^marker>\tag unimportant\ 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 \Factoring out a zero according to its order\ +subsection\<^marker>\tag unimportant\ \Factoring out a zero according to its order\ lemma holomorphic_factor_order_of_zero: assumes holf: "f holomorphic_on S" @@ -1021,7 +1021,7 @@ qed qed -subsection%unimportant \Entire proper functions are precisely the non-trivial polynomials\ +subsection\<^marker>\tag unimportant\ \Entire proper functions are precisely the non-trivial polynomials\ lemma proper_map_polyfun: fixes c :: "nat \ 'a::{real_normed_div_algebra,heine_borel}" @@ -2163,7 +2163,7 @@ text\Wenda Li and LC Paulson (2016). A Formal Proof of Cauchy's Residue Theorem. Interactive Theorem Proving\ -definition%important residue :: "(complex \ complex) \ complex \ complex" where +definition\<^marker>\tag important\ residue :: "(complex \ complex) \ complex \ complex" where "residue f z = (SOME int. \e>0. \\>0. \ (f has_contour_integral 2*pi* \ *int) (circlepath z \))" @@ -2967,7 +2967,7 @@ subsection \Non-essential singular points\ -definition%important is_pole :: +definition\<^marker>\tag important\ is_pole :: "('a::topological_space \ 'b::real_normed_vector) \ 'a \ bool" where "is_pole f a = (LIM x (at a). f x :> at_infinity)" @@ -3695,12 +3695,12 @@ subsubsection \The order of non-essential singularities (i.e. removable singularities or poles)\ -definition%important zorder :: "(complex \ complex) \ complex \ int" where +definition\<^marker>\tag important\ zorder :: "(complex \ complex) \ complex \ int" where "zorder f z = (THE n. (\h r. r>0 \ h holomorphic_on cball z r \ h z\0 \ (\w\cball z r - {z}. f w = h w * (w-z) powr (of_int n) \ h w \0)))" -definition%important zor_poly +definition\<^marker>\tag important\ zor_poly ::"[complex \ complex, complex] \ complex \ complex" where "zor_poly f z = (SOME h. \r. r > 0 \ h holomorphic_on cball z r \ h z \ 0 \ (\w\cball z r - {z}. f w = h w * (w - z) powr (zorder f z) diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Connected.thy --- 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 \Connectedness\ +subsection\<^marker>\tag unimportant\ \Connectedness\ lemma connected_local: "connected S \ @@ -68,7 +68,7 @@ subsection \Connected components, considered as a connectedness relation or a set\ -definition%important "connected_component s x y \ \t. connected t \ t \ s \ x \ t \ y \ t" +definition\<^marker>\tag important\ "connected_component s x y \ \t. connected t \ t \ s \ x \ t \ y \ t" abbreviation "connected_component_set s x \ Collect (connected_component s x)" @@ -268,7 +268,7 @@ subsection \The set of connected components of a set\ -definition%important components:: "'a::topological_space set \ 'a set set" +definition\<^marker>\tag important\ components:: "'a::topological_space set \ 'a set set" where "components s \ connected_component_set s ` s" lemma components_iff: "s \ components u \ (\x. x \ u \ s = connected_component_set u x)" @@ -427,7 +427,7 @@ using closedin_connected_component componentsE by blast -subsection%unimportant \Proving a function is constant on a connected set +subsection\<^marker>\tag unimportant\ \Proving a function is constant on a connected set by proving that a level set is open\ lemma continuous_levelset_openin_cases: @@ -458,7 +458,7 @@ by fast -subsection%unimportant \Preservation of Connectedness\ +subsection\<^marker>\tag unimportant\ \Preservation of Connectedness\ lemma homeomorphic_connectedness: assumes "s homeomorphic t" @@ -688,7 +688,7 @@ qed -subsection%unimportant\Constancy of a function from a connected set into a finite, disconnected or discrete set\ +subsection\<^marker>\tag unimportant\\Constancy of a function from a connected set into a finite, disconnected or discrete set\ text\Still missing: versions for a set that is smaller than R, or countable.\ diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Continuous_Extension.thy --- 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 @@ "\x. x \ U \ norm(g x) \ 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>\tag unimportant\ Tietze_closed_interval: fixes f :: "'a::{metric_space,second_countable_topology} \ '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>\tag unimportant\ Tietze_closed_interval_1: fixes f :: "'a::{metric_space,second_countable_topology} \ 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>\tag unimportant\ Tietze_open_interval: fixes f :: "'a::{metric_space,second_countable_topology} \ '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>\tag unimportant\ Tietze_open_interval_1: fixes f :: "'a::{metric_space,second_countable_topology} \ 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>\tag unimportant\ Tietze_unbounded: fixes f :: "'a::{metric_space,second_countable_topology} \ 'b::real_inner" assumes "continuous_on S f" and "closedin (top_of_set U) S" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Continuum_Not_Denumerable.thy --- 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 \Abstract\ +subsection\<^marker>\tag unimportant\ \Abstract\ text \ The following document presents a proof that the Continuum is uncountable. @@ -32,7 +32,7 @@ range of \f\ by generating a sequence of closed intervals then using the Nested Interval Property. \ -text%important \%whitespace\ +text\<^marker>\tag important\ \%whitespace\ theorem real_non_denum: "\f :: nat \ real. surj f" proof assume "\f::nat \ real. surj f" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Convex.thy --- 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 \Convexity\ -definition%important convex :: "'a::real_vector set \ bool" +definition\<^marker>\tag important\ convex :: "'a::real_vector set \ bool" where "convex s \ (\x\s. \y\s. \u\0. \v\0. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ s)" lemma convexI: @@ -189,7 +189,7 @@ by (simp add: convex_def scaleR_conv_of_real) -subsection%unimportant \Explicit expressions for convexity in terms of arbitrary sums\ +subsection\<^marker>\tag unimportant\ \Explicit expressions for convexity in terms of arbitrary sums\ lemma convex_sum: fixes C :: "'a::real_vector set" @@ -342,7 +342,7 @@ subsection \Functions that are convex on a set\ -definition%important convex_on :: "'a::real_vector set \ ('a \ real) \ bool" +definition\<^marker>\tag important\ convex_on :: "'a::real_vector set \ ('a \ real) \ bool" where "convex_on s f \ (\x\s. \y\s. \u\0. \v\0. u + v = 1 \ f (u *\<^sub>R x + v *\<^sub>R y) \ u * f x + v * f y)" @@ -463,7 +463,7 @@ qed -subsection%unimportant \Arithmetic operations on sets preserve convexity\ +subsection\<^marker>\tag unimportant\ \Arithmetic operations on sets preserve convexity\ lemma convex_linear_image: assumes "linear f" @@ -929,7 +929,7 @@ qed -subsection%unimportant \Convexity of real functions\ +subsection\<^marker>\tag unimportant\ \Convexity of real functions\ lemma convex_on_realI: assumes "connected A" @@ -1111,7 +1111,7 @@ subsection \Affine set and affine hull\ -definition%important affine :: "'a::real_vector set \ bool" +definition\<^marker>\tag important\ affine :: "'a::real_vector set \ bool" where "affine s \ (\x\s. \y\s. \u v. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ s)" lemma affine_alt: "affine s \ (\x\s. \y\s. \u::real. (1 - u) *\<^sub>R x + u *\<^sub>R y \ s)" @@ -1149,7 +1149,7 @@ by (simp add: affine_def algebra_simps) (metis distrib_right mult.left_neutral) -subsubsection%unimportant \Some explicit formulations\ +subsubsection\<^marker>\tag unimportant\ \Some explicit formulations\ text "Formalized by Lars Schewe." @@ -1305,7 +1305,7 @@ qed -subsubsection%unimportant \Stepping theorems and hence small special cases\ +subsubsection\<^marker>\tag unimportant\ \Stepping theorems and hence small special cases\ 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 \Some relations between affine hull and subspaces\ +subsubsection\<^marker>\tag unimportant\ \Some relations between affine hull and subspaces\ lemma affine_hull_insert_subset_span: "affine hull (insert a S) \ {a + v| v . v \ span {x - a | x . x \ S}}" @@ -1473,7 +1473,7 @@ using affine_hull_insert_span[of a "S - {a}", unfolded insert_Diff[OF assms]] by auto -subsubsection%unimportant \Parallel affine sets\ +subsubsection\<^marker>\tag unimportant\ \Parallel affine sets\ definition affine_parallel :: "'a::real_vector set \ 'a::real_vector set \ bool" where "affine_parallel S T \ (\a. T = (\x. a + x) ` S)" @@ -1573,7 +1573,7 @@ unfolding subspace_def affine_def by auto -subsubsection%unimportant \Subspace parallel to an affine set\ +subsubsection\<^marker>\tag unimportant\ \Subspace parallel to an affine set\ lemma subspace_affine: "subspace S \ affine S \ 0 \ S" proof - @@ -1703,7 +1703,7 @@ subsection \Cones\ -definition%important cone :: "'a::real_vector set \ bool" +definition\<^marker>\tag important\ cone :: "'a::real_vector set \ bool" where "cone s \ (\x\s. \c\0. c *\<^sub>R x \ s)" lemma cone_empty[intro, simp]: "cone {}" @@ -1869,7 +1869,7 @@ text "Formalized by Lars Schewe." -definition%important affine_dependent :: "'a::real_vector set \ bool" +definition\<^marker>\tag important\ affine_dependent :: "'a::real_vector set \ bool" where "affine_dependent s \ (\x\s. x \ affine hull (s - {x}))" lemma affine_dependent_subset: @@ -1946,7 +1946,7 @@ qed -subsection%unimportant \Connectedness of convex sets\ +subsection\<^marker>\tag unimportant\ \Connectedness of convex sets\ lemma connectedD: "connected S \ open A \ open B \ S \ A \ B \ A \ B \ S = {} \ A \ S = {} \ B \ S = {}" @@ -2003,7 +2003,7 @@ lemma convex_hull_eq: "convex hull s = s \ convex s" by (metis convex_convex_hull hull_same) -subsubsection%unimportant \Convex hull is "preserved" by a linear function\ +subsubsection\<^marker>\tag unimportant\ \Convex hull is "preserved" by a linear function\ lemma convex_hull_linear_image: assumes f: "linear f" @@ -2061,7 +2061,7 @@ qed -subsubsection%unimportant \Stepping theorems for convex hulls of finite sets\ +subsubsection\<^marker>\tag unimportant\ \Stepping theorems for convex hulls of finite sets\ 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 \Explicit expression for convex hull\ +subsubsection\<^marker>\tag unimportant\ \Explicit expression for convex hull\ proposition convex_hull_indexed: fixes S :: "'a::real_vector set" @@ -2262,7 +2262,7 @@ qed (use assms in \auto simp: convex_explicit\) -subsubsection%unimportant \Another formulation\ +subsubsection\<^marker>\tag unimportant\ \Another formulation\ text "Formalized by Lars Schewe." @@ -2362,7 +2362,7 @@ qed -subsubsection%unimportant \A stepping theorem for that expansion\ +subsubsection\<^marker>\tag unimportant\ \A stepping theorem for that expansion\ lemma convex_hull_finite_step: fixes S :: "'a::real_vector set" @@ -2419,7 +2419,7 @@ qed -subsubsection%unimportant \Hence some special cases\ +subsubsection\<^marker>\tag unimportant\ \Hence some special cases\ lemma convex_hull_2: "convex hull {a,b} = {u *\<^sub>R a + v *\<^sub>R b | u v. 0 \ u \ 0 \ v \ u + v = 1}" @@ -2491,7 +2491,7 @@ qed -subsection%unimportant \Relations among closure notions and corresponding hulls\ +subsection\<^marker>\tag unimportant\ \Relations among closure notions and corresponding hulls\ lemma affine_imp_convex: "affine s \ convex s" unfolding affine_def convex_def by auto @@ -2629,7 +2629,7 @@ qed -subsection%unimportant \Some Properties of Affine Dependent Sets\ +subsection\<^marker>\tag unimportant\ \Some Properties of Affine Dependent Sets\ lemma affine_independent_0 [simp]: "\ affine_dependent {}" by (simp add: affine_dependent_def) @@ -2868,7 +2868,7 @@ subsection \Affine Dimension of a Set\ -definition%important aff_dim :: "('a::euclidean_space) set \ int" +definition\<^marker>\tag important\ aff_dim :: "('a::euclidean_space) set \ int" where "aff_dim V = (SOME d :: int. \B. affine hull B = affine hull V \ \ affine_dependent B \ of_nat (card B) = d + 1)" @@ -3656,7 +3656,7 @@ using hull_mono[OF \s\p\] by auto qed -subsection%unimportant\Some Properties of subset of standard basis\ +subsection\<^marker>\tag unimportant\\Some Properties of subset of standard basis\ lemma affine_hull_substd_basis: assumes "d \ Basis" @@ -3673,7 +3673,7 @@ by (metis Int_absorb1 Int_absorb2 convex_hull_subset_affine_hull hull_hull hull_mono hull_subset) -subsection%unimportant \Moving and scaling convex hulls\ +subsection\<^marker>\tag unimportant\ \Moving and scaling convex hulls\ 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 \Convexity of cone hulls\ +subsection\<^marker>\tag unimportant\ \Convexity of cone hulls\ lemma convex_cone_hull: assumes "convex S" @@ -4000,7 +4000,7 @@ subsection \Epigraphs of convex functions\ -definition%important "epigraph S (f :: _ \ real) = {xy. fst xy \ S \ f (fst xy) \ snd xy}" +definition\<^marker>\tag important\ "epigraph S (f :: _ \ real) = {xy. fst xy \ S \ f (fst xy) \ snd xy}" lemma mem_epigraph: "(x, y) \ epigraph S f \ x \ S \ f x \ y" unfolding epigraph_def by auto @@ -4036,7 +4036,7 @@ by (simp add: convex_epigraph) -subsubsection%unimportant \Use this to derive general bound property of convex function\ +subsubsection\<^marker>\tag unimportant\ \Use this to derive general bound property of convex function\ lemma convex_on: assumes "convex S" @@ -4060,7 +4060,7 @@ using assms[unfolded convex] apply auto done -subsection%unimportant \A bound within a convex hull\ +subsection\<^marker>\tag unimportant\ \A bound within a convex hull\ lemma convex_on_convex_hull_bound: assumes "convex_on (convex hull s) f" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Convex_Euclidean_Space.thy --- 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 \Topological Properties of Convex Sets and Functions\ +subsection\<^marker>\tag unimportant\ \Topological Properties of Convex Sets and Functions\ lemma convex_supp_sum: assumes "convex S" and 1: "supp_sum u I = 1" @@ -318,7 +318,7 @@ subsection \Relative interior of a set\ -definition%important "rel_interior S = +definition\<^marker>\tag important\ "rel_interior S = {x. \T. openin (top_of_set (affine hull S)) T \ x \ T \ T \ S}" lemma rel_interior_mono: @@ -649,7 +649,7 @@ subsubsection \Relative open sets\ -definition%important "rel_open S \ rel_interior S = S" +definition\<^marker>\tag important\ "rel_open S \ rel_interior S = S" lemma rel_open: "rel_open S \ openin (top_of_set (affine hull S)) S" unfolding rel_open_def rel_interior_def @@ -848,7 +848,7 @@ qed -subsubsection%unimportant\Relative interior preserves under linear transformations\ +subsubsection\<^marker>\tag unimportant\\Relative interior preserves under linear transformations\ lemma rel_interior_translation_aux: fixes a :: "'n::euclidean_space" @@ -1005,7 +1005,7 @@ by auto -subsection%unimportant \Openness and compactness are preserved by convex hull operation\ +subsection\<^marker>\tag unimportant\ \Openness and compactness are preserved by convex hull operation\ lemma open_convex_hull[intro]: fixes S :: "'a::real_normed_vector set" @@ -1237,7 +1237,7 @@ qed -subsection%unimportant \Extremal points of a simplex are some vertices\ +subsection\<^marker>\tag unimportant\ \Extremal points of a simplex are some vertices\ lemma dist_increases_online: fixes a b d :: "'a::real_inner" @@ -1432,7 +1432,7 @@ subsection \Closest point of a convex set is unique, with a continuous projection\ -definition%important closest_point :: "'a::{real_inner,heine_borel} set \ 'a \ 'a" +definition\<^marker>\tag important\ closest_point :: "'a::{real_inner,heine_borel} set \ 'a \ 'a" where "closest_point S a = (SOME x. x \ S \ (\y\S. dist a x \ dist a y))" lemma closest_point_exists: @@ -1622,7 +1622,7 @@ qed -subsubsection%unimportant \Various point-to-set separating/supporting hyperplane theorems\ +subsubsection\<^marker>\tag unimportant\ \Various point-to-set separating/supporting hyperplane theorems\ lemma supporting_hyperplane_closed_point: fixes z :: "'a::{real_inner,heine_borel}" @@ -1707,7 +1707,7 @@ qed -subsubsection%unimportant \Now set-to-set for closed/compact sets\ +subsubsection\<^marker>\tag unimportant\ \Now set-to-set for closed/compact sets\ lemma separating_hyperplane_closed_compact: fixes S :: "'a::euclidean_space set" @@ -1797,7 +1797,7 @@ qed -subsubsection%unimportant \General case without assuming closure and getting non-strict separation\ +subsubsection\<^marker>\tag unimportant\ \General case without assuming closure and getting non-strict separation\ lemma separating_hyperplane_set_0: assumes "convex S" "(0::'a::euclidean_space) \ S" @@ -1855,7 +1855,7 @@ qed -subsection%unimportant \More convexity generalities\ +subsection\<^marker>\tag unimportant\ \More convexity generalities\ 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 \Convex set as intersection of halfspaces\ +subsection\<^marker>\tag unimportant\ \Convex set as intersection of halfspaces\ lemma convex_halfspace_intersection: fixes s :: "('a::euclidean_space) set" @@ -1922,7 +1922,7 @@ qed auto -subsection%unimportant \Convexity of general and special intervals\ +subsection\<^marker>\tag unimportant\ \Convexity of general and special intervals\ lemma is_interval_convex: fixes S :: "'a::euclidean_space set" @@ -2007,7 +2007,7 @@ "\connected S; aff_dim S \ 0; a \ S\ \ a islimpt S" using aff_dim_sing connected_imp_perfect by blast -subsection%unimportant \On \real\, \is_interval\, \convex\ and \connected\ are all equivalent\ +subsection\<^marker>\tag unimportant\ \On \real\, \is_interval\, \convex\ and \connected\ are all equivalent\ lemma mem_is_interval_1_I: fixes a b c::real @@ -2099,7 +2099,7 @@ and is_interval_co: "is_interval {r..Another intermediate value theorem formulation\ +subsection\<^marker>\tag unimportant\ \Another intermediate value theorem formulation\ lemma ivt_increasing_component_on_1: fixes f :: "real \ 'a::euclidean_space" @@ -2144,7 +2144,7 @@ by (rule ivt_decreasing_component_on_1) (auto simp: continuous_at_imp_continuous_on) -subsection%unimportant \A bound within an interval\ +subsection\<^marker>\tag unimportant\ \A bound within an interval\ lemma convex_hull_eq_real_cbox: fixes x y :: real assumes "x \ y" @@ -2235,7 +2235,7 @@ by (rule_tac that[of "(\y. x - ?d + (2 * d) *\<^sub>R y)` S"]) (auto simp: convex_hull_affinity *) qed -subsection%unimportant\Representation of any interval as a finite convex hull\ +subsection\<^marker>\tag unimportant\\Representation of any interval as a finite convex hull\ lemma image_stretch_interval: "(\x. \k\Basis. (m k * (x\k)) *\<^sub>R k) ` cbox a (b::'a::euclidean_space) = @@ -2312,7 +2312,7 @@ qed -subsection%unimportant \Bounded convex function on open set is continuous\ +subsection\<^marker>\tag unimportant\ \Bounded convex function on open set is continuous\ lemma convex_on_bounded_continuous: fixes S :: "('a::real_normed_vector) set" @@ -2405,7 +2405,7 @@ qed -subsection%unimportant \Upper bound on a ball implies upper and lower bounds\ +subsection\<^marker>\tag unimportant\ \Upper bound on a ball implies upper and lower bounds\ lemma convex_bounds_lemma: fixes x :: "'a::real_normed_vector" @@ -2439,7 +2439,7 @@ qed -subsubsection%unimportant \Hence a convex function on an open set is continuous\ +subsubsection\<^marker>\tag unimportant\ \Hence a convex function on an open set is continuous\ lemma real_of_nat_ge_one_iff: "1 \ real (n::nat) \ 1 \ n" by auto diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Cross3.thy --- 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 \\locally disable syntax for set product, to avoid warnings\ -definition%important cross3 :: "[real^3, real^3] \ real^3" (infixr "\" 80) +definition\<^marker>\tag important\ cross3 :: "[real^3, real^3] \ real^3" (infixr "\" 80) where "a \ b \ vector [a$2 * b$3 - a$3 * b$2, a$3 * b$1 - a$1 * b$3, diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Derivative.thy --- 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 \Derivative with composed bilinear function\ +subsection\<^marker>\tag unimportant\ \Derivative with composed bilinear function\ text \More explicit epsilon-delta forms.\ @@ -92,7 +92,7 @@ subsection \Differentiability\ -definition%important +definition\<^marker>\tag important\ differentiable_on :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a set \ bool" (infix "differentiable'_on" 50) where "f differentiable_on s \ (\x\s. f differentiable (at x within s))" @@ -113,7 +113,7 @@ "(\x. x \ s \ f differentiable at x) \ f differentiable_on s" by (metis differentiable_at_withinI differentiable_on_def) -corollary%unimportant differentiable_iff_scaleR: +corollary\<^marker>\tag unimportant\ differentiable_iff_scaleR: fixes f :: "real \ 'a::real_normed_vector" shows "f differentiable F \ (\d. (f has_derivative (\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 \Composition rules stated just for differentiability\ +subsection\<^marker>\tag unimportant\ \Composition rules stated just for differentiability\ lemma differentiable_chain_at: "f differentiable (at x) \ @@ -342,7 +342,7 @@ subsection \Uniqueness of derivative\ -text%important \ +text\<^marker>\tag important\ \ The general result is a bit messy because we need approachability of the limit point from any direction. But OK for nontrivial intervals etc. \ @@ -2195,7 +2195,7 @@ subsection \Field differentiability\ -definition%important field_differentiable :: "['a \ 'a::real_normed_field, 'a filter] \ bool" +definition\<^marker>\tag important\ field_differentiable :: "['a \ 'a::real_normed_field, 'a filter] \ bool" (infixr "(field'_differentiable)" 50) where "f field_differentiable F \ \f'. (f has_field_derivative f') F" @@ -2390,7 +2390,7 @@ subsection \Field derivative\ -definition%important deriv :: "('a \ 'a::real_normed_field) \ 'a \ 'a" where +definition\<^marker>\tag important\ deriv :: "('a \ 'a::real_normed_field) \ 'a \ 'a" where "deriv f x \ SOME D. DERIV f x :> D" lemma DERIV_imp_deriv: "DERIV f x :> f' \ deriv f x = f'" @@ -2701,7 +2701,7 @@ qed -subsection%unimportant \Differentiable case distinction\ +subsection\<^marker>\tag unimportant\ \Differentiable case distinction\ lemma has_derivative_within_If_eq: "((\x. if P x then f x else g x) has_derivative f') (at x within S) = diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Determinants.thy --- 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 \Trace\ -definition%important trace :: "'a::semiring_1^'n^'n \ 'a" +definition\<^marker>\tag important\ trace :: "'a::semiring_1^'n^'n \ 'a" where "trace A = sum (\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 \Definition of determinant\ +subsubsection\<^marker>\tag important\ \Definition of determinant\ -definition%important det:: "'a::comm_ring_1^'n^'n \ 'a" where +definition\<^marker>\tag important\ det:: "'a::comm_ring_1^'n^'n \ 'a" where "det A = sum (\p. of_int (sign p) * prod (\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 \Invertibility of matrices and corresponding linear functions\ +subsubsection\<^marker>\tag important\ \Invertibility of matrices and corresponding linear functions\ lemma matrix_left_invertible_gen: fixes f :: "'a::field^'m \ 'a::field^'n" @@ -994,12 +994,12 @@ proposition orthogonal_transformation_det [simp]: fixes f :: "real^'n \ real^'n" shows "orthogonal_transformation f \ \det (matrix f)\ = 1" - using%unimportant det_orthogonal_matrix orthogonal_transformation_matrix by fastforce + using det_orthogonal_matrix orthogonal_transformation_matrix by fastforce subsection \Rotation, reflection, rotoinversion\ -definition%important "rotation_matrix Q \ orthogonal_matrix Q \ det Q = 1" -definition%important "rotoinversion_matrix Q \ orthogonal_matrix Q \ det Q = - 1" +definition\<^marker>\tag important\ "rotation_matrix Q \ orthogonal_matrix Q \ det Q = 1" +definition\<^marker>\tag important\ "rotoinversion_matrix Q \ orthogonal_matrix Q \ det Q = - 1" lemma orthogonal_rotation_or_rotoinversion: fixes Q :: "'a::linordered_idom^'n^'n" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Elementary_Metric_Spaces.thy --- 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 and closed balls\ -definition%important ball :: "'a::metric_space \ real \ 'a set" +definition\<^marker>\tag important\ ball :: "'a::metric_space \ real \ 'a set" where "ball x e = {y. dist x y < e}" -definition%important cball :: "'a::metric_space \ real \ 'a set" +definition\<^marker>\tag important\ cball :: "'a::metric_space \ real \ 'a set" where "cball x e = {y. dist x y \ e}" -definition%important sphere :: "'a::metric_space \ real \ 'a set" +definition\<^marker>\tag important\ sphere :: "'a::metric_space \ real \ 'a set" where "sphere x e = {y. dist x y = e}" lemma mem_ball [simp]: "y \ ball x e \ dist x y < e" @@ -717,7 +717,7 @@ subsection \Boundedness\ (* FIXME: This has to be unified with BSEQ!! *) -definition%important (in metric_space) bounded :: "'a set \ bool" +definition\<^marker>\tag important\ (in metric_space) bounded :: "'a set \ bool" where "bounded S \ (\x e. \y\S. dist x y \ e)" lemma bounded_subset_cball: "bounded S \ (\e x. S \ cball x e \ 0 \ e)" @@ -1198,7 +1198,7 @@ subsection \The diameter of a set\ -definition%important diameter :: "'a::metric_space set \ real" where +definition\<^marker>\tag important\ diameter :: "'a::metric_space set \ real" where "diameter S = (if S = {} then 0 else SUP (x,y)\S\S. dist x y)" lemma diameter_empty [simp]: "diameter{} = 0" @@ -1464,8 +1464,8 @@ shows "compact(closure S) \ 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>\tag important\ real :: heine_borel +proof fix f :: "nat \ real" assume f: "bounded (range f)" obtain r :: "nat \ nat" where r: "strict_mono r" "monoseq (f \ 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>\tag important\ prod :: (heine_borel, heine_borel) heine_borel +proof fix f :: "nat \ 'a \ 'b" assume f: "bounded (range f)" then have "bounded (fst ` range f)" @@ -1844,7 +1844,7 @@ by auto -subsection%unimportant\ Finite intersection property\ +subsection\<^marker>\tag unimportant\\ Finite intersection property\ text\Also developed in HOL's toplogical spaces theory, but the Heine-Borel type class isn't available there.\ @@ -1961,7 +1961,7 @@ subsection \Infimum Distance\ -definition%important "infdist x A = (if A = {} then 0 else INF a\A. dist x a)" +definition\<^marker>\tag important\ "infdist x A = (if A = {} then 0 else INF a\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\ Theorems relating continuity and uniform continuity to closures\ +subsection\<^marker>\tag unimportant\\ Theorems relating continuity and uniform continuity to closures\ lemma continuous_on_closure: "continuous_on (closure S) f \ @@ -2860,7 +2860,7 @@ then show ?thesis .. qed -subsection%unimportant \Making a continuous function avoid some value in a neighbourhood\ +subsection\<^marker>\tag unimportant\ \Making a continuous function avoid some value in a neighbourhood\ lemma continuous_within_avoid: fixes f :: "'a::metric_space \ 'b::t1_space" @@ -3041,7 +3041,7 @@ subsection\The infimum of the distance between two sets\ -definition%important setdist :: "'a::metric_space set \ 'a set \ real" where +definition\<^marker>\tag important\ setdist :: "'a::metric_space set \ 'a set \ real" where "setdist s t \ (if s = {} \ t = {} then 0 else Inf {dist x y| x y. x \ s \ y \ t})" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Elementary_Normed_Spaces.thy --- 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 \Various Lemmas Combining Imports\ +subsection\<^marker>\tag unimportant\ \Various Lemmas Combining Imports\ lemma countable_PiE: "finite I \ (\i. i \ I \ countable (F i)) \ countable (Pi\<^sub>E I F)" @@ -428,7 +428,7 @@ done -subsection%unimportant \Various Lemmas on Normed Algebras\ +subsection\<^marker>\tag unimportant\ \Various Lemmas on Normed Algebras\ 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 \Limit Point of Filter\ +subsection\<^marker>\tag unimportant\ \Limit Point of Filter\ lemma netlimit_at_vector: fixes a :: "'a::real_normed_vector" @@ -775,7 +775,7 @@ shows "bounded (- S) \ \ bounded S" using bounded_Un [of S "-S"] by (simp add: sup_compl_top) -subsection%unimportant\Relations among convergence and absolute convergence for power series\ +subsection\<^marker>\tag unimportant\\Relations among convergence and absolute convergence for power series\ lemma summable_imp_bounded: fixes f :: "nat \ 'a::real_normed_vector" @@ -1024,7 +1024,7 @@ subsection \Continuity\ -subsubsection%unimportant \Structural rules for uniform continuity\ +subsubsection\<^marker>\tag unimportant\ \Structural rules for uniform continuity\ lemma (in bounded_linear) uniformly_continuous_on[continuous_intros]: fixes g :: "_::metric_space \ _" @@ -1103,7 +1103,7 @@ by (simp add: fun_Compl_def uniformly_continuous_on_minus) -subsection%unimportant \Topological properties of linear functions\ +subsection\<^marker>\tag unimportant\ \Topological properties of linear functions\ lemma linear_lim_0: assumes "bounded_linear f" @@ -1131,7 +1131,7 @@ "bounded_linear f \ continuous_on s f" using continuous_at_imp_continuous_on[of s f] using linear_continuous_at[of f] by auto -subsection%unimportant \Arithmetic Preserves Topological Properties\ +subsection\<^marker>\tag unimportant\ \Arithmetic Preserves Topological Properties\ 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\Homeomorphisms\ +subsection\<^marker>\tag unimportant\\Homeomorphisms\ 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 \Discrete\ +subsection\<^marker>\tag unimportant\ \Discrete\ lemma finite_implies_discrete: fixes S :: "'a::topological_space set" @@ -1627,7 +1627,7 @@ qed -subsection%unimportant \Completeness of "Isometry" (up to constant bounds)\ +subsection\<^marker>\tag unimportant\ \Completeness of "Isometry" (up to constant bounds)\ lemma cauchy_isometric:\ \TODO: rename lemma to \Cauchy_isometric\\ assumes e: "e > 0" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Elementary_Topology.thy --- 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 \Archimedean properties and useful consequences\ +subsubsection\<^marker>\tag unimportant\ \Archimedean properties and useful consequences\ text\Bernoulli's inequality\ proposition Bernoulli_inequality: @@ -111,7 +111,7 @@ apply (metis Suc_pred of_nat_Suc) done -subsubsection%unimportant \Affine transformations of intervals\ +subsubsection\<^marker>\tag unimportant\ \Affine transformations of intervals\ lemma real_affinity_le: "0 < m \ m * x + c \ y \ x \ inverse m * y + - (c / m)" for m :: "'a::linordered_field" @@ -144,7 +144,7 @@ context topological_space begin -definition%important "topological_basis B \ +definition\<^marker>\tag important\ "topological_basis B \ (\b\B. open b) \ (\x. open x \ (\B'. B' \ B \ \B' = x))" lemma topological_basis: @@ -270,7 +270,7 @@ subsection \Countable Basis\ -locale%important countable_basis = topological_space p for p::"'a set \ bool" + +locale\<^marker>\tag important\ countable_basis = topological_space p for p::"'a set \ bool" + fixes B :: "'a set set" assumes is_basis: "topological_basis B" and countable_basis: "countable B" @@ -666,7 +666,7 @@ subsection \Limit Points\ -definition%important (in topological_space) islimpt:: "'a \ 'a set \ bool" (infixr "islimpt" 60) +definition\<^marker>\tag important\ (in topological_space) islimpt:: "'a \ 'a set \ bool" (infixr "islimpt" 60) where "x islimpt S \ (\T. x\T \ open T \ (\y\S. y\T \ y\x))" lemma islimptI: @@ -866,7 +866,7 @@ subsection \Interior of a Set\ -definition%important interior :: "('a::topological_space) set \ 'a set" where +definition\<^marker>\tag important\ interior :: "('a::topological_space) set \ 'a set" where "interior S = \{T. open T \ T \ S}" lemma interiorI [intro?]: @@ -1045,7 +1045,7 @@ subsection \Closure of a Set\ -definition%important closure :: "('a::topological_space) set \ 'a set" where +definition\<^marker>\tag important\ closure :: "('a::topological_space) set \ 'a set" where "closure S = S \ {x . x islimpt S}" lemma interior_closure: "interior S = - (closure (- S))" @@ -1175,7 +1175,7 @@ subsection \Frontier (also known as boundary)\ -definition%important frontier :: "('a::topological_space) set \ 'a set" where +definition\<^marker>\tag important\ frontier :: "('a::topological_space) set \ 'a set" where "frontier S = closure S - interior S" lemma frontier_closed [iff]: "closed (frontier S)" @@ -1253,7 +1253,7 @@ qed -subsection%unimportant \Filters and the ``eventually true'' quantifier\ +subsection\<^marker>\tag unimportant\ \Filters and the ``eventually true'' quantifier\ text \Identify Trivial limits, where we can't approach arbitrarily closely.\ @@ -1846,7 +1846,7 @@ with \U \ \A = {}\ show False by auto qed -definition%important countably_compact :: "('a::topological_space) set \ bool" where +definition\<^marker>\tag important\ countably_compact :: "('a::topological_space) set \ bool" where "countably_compact U \ (\A. countable A \ (\a\A. open a) \ U \ \A \ (\T\A. finite T \ U \ \T))" @@ -1916,7 +1916,7 @@ subsubsection\Sequential compactness\ -definition%important seq_compact :: "'a::topological_space set \ bool" where +definition\<^marker>\tag important\ seq_compact :: "'a::topological_space set \ bool" where "seq_compact S \ (\f. (\n. f n \ S) \ (\l\S. \r::nat\nat. strict_mono r \ ((f \ r) \ l) sequentially))" @@ -2184,7 +2184,7 @@ by (rule countable_acc_point_imp_seq_compact) (metis islimpt_eq_acc_point) -subsection%unimportant \Cartesian products\ +subsection\<^marker>\tag unimportant\ \Cartesian products\ lemma seq_compact_Times: "seq_compact s \ seq_compact t \ seq_compact (s \ t)" unfolding seq_compact_def @@ -2457,7 +2457,7 @@ subsection \Homeomorphisms\ -definition%important "homeomorphism s t f g \ +definition\<^marker>\tag important\ "homeomorphism s t f g \ (\x\s. (g(f x) = x)) \ (f ` s = t) \ continuous_on s f \ (\y\t. (f(g y) = y)) \ (g ` t = s) \ 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 \ 'b::topological_space set \ bool" +definition\<^marker>\tag important\ homeomorphic :: "'a::topological_space set \ 'b::topological_space set \ bool" (infixr "homeomorphic" 60) where "s homeomorphic t \ (\f g. homeomorphism s t f g)" @@ -2661,7 +2661,7 @@ by (metis compact_continuous_image) -subsection%unimportant \On Linorder Topologies\ +subsection\<^marker>\tag unimportant\ \On Linorder Topologies\ lemma islimpt_greaterThanLessThan1: fixes a b::"'a::{linorder_topology, dense_order}" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Embed_Measure.thy --- 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 \RealVal\ for real numbers. Then \embed_measure\ allows us to lift a measure on real numbers to the appropriate subset of that algebraic datatype. \ -definition%important embed_measure :: "'a measure \ ('a \ 'b) \ 'b measure" where +definition\<^marker>\tag important\ embed_measure :: "'a measure \ ('a \ 'b) \ 'b measure" where "embed_measure M f = measure_of (f ` space M) {f ` A |A. A \ sets M} (\A. emeasure M (f -` A \ space M))" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Euclidean_Space.thy --- 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 \Interlude: Some properties of real sets\ +subsection\<^marker>\tag unimportant\ \Interlude: Some properties of real sets\ lemma seq_mono_lemma: assumes "\(n::nat) \ m. (d n :: real) < e n" @@ -211,7 +211,7 @@ qed -subsection%unimportant \Subclass relationships\ +subsection\<^marker>\tag unimportant\ \Subclass relationships\ instance euclidean_space \ perfect_space proof @@ -234,7 +234,7 @@ subsection \Class instances\ -subsubsection%unimportant \Type \<^typ>\real\\ +subsubsection\<^marker>\tag unimportant\ \Type \<^typ>\real\\ instantiation real :: euclidean_space begin @@ -250,7 +250,7 @@ lemma DIM_real[simp]: "DIM(real) = 1" by simp -subsubsection%unimportant \Type \<^typ>\complex\\ +subsubsection\<^marker>\tag unimportant\ \Type \<^typ>\complex\\ instantiation complex :: euclidean_space begin @@ -271,7 +271,7 @@ lemma complex_Basis_i [iff]: "\ \ Basis" by (simp add: Basis_complex_def) -subsubsection%unimportant \Type \<^typ>\'a \ 'b\\ +subsubsection\<^marker>\tag unimportant\ \Type \<^typ>\'a \ 'b\\ instantiation prod :: (real_inner, real_inner) real_inner begin diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Extended_Real_Limits.thy --- 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 \ 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>\tag unimportant\ enat :: second_countable_topology proof show "\B::enat set set. countable B \ 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>\tag unimportant\ ereal :: second_countable_topology proof (standard, intro exI conjI) let ?B = "(\r\\. {{..< r}, {r <..}} :: ereal set set)" show "countable ?B" @@ -391,7 +391,7 @@ qed -subsubsection%important \Continuity of addition\ +subsubsection\<^marker>\tag important\ \Continuity of addition\ text \The next few lemmas remove an unnecessary assumption in \tendsto_add_ereal\, culminating in \tendsto_add_ereal_general\ 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 \Continuity of multiplication\ +subsubsection\<^marker>\tag important\ \Continuity of multiplication\ text \In the same way as for addition, we prove that the multiplication is continuous on ereal times ereal, except at \(\, 0)\ and \(-\, 0)\ and \(0, \)\ and \(0, -\)\, @@ -659,7 +659,7 @@ by (cases "c = 0", auto simp add: assms tendsto_mult_ereal) -subsubsection%important \Continuity of division\ +subsubsection\<^marker>\tag important\ \Continuity of division\ lemma tendsto_inverse_ereal_PInf: fixes u::"_ \ ereal" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/FPS_Convergence.thy --- 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 \Balls with extended real radius\ +subsection\<^marker>\tag unimportant\ \Balls with extended real radius\ text \ The following is a variant of \<^const>\ball\ that also allows an infinite radius. @@ -55,13 +55,13 @@ subsection \Basic properties of convergent power series\ -definition%important fps_conv_radius :: "'a :: {banach, real_normed_div_algebra} fps \ ereal" where +definition\<^marker>\tag important\ fps_conv_radius :: "'a :: {banach, real_normed_div_algebra} fps \ ereal" where "fps_conv_radius f = conv_radius (fps_nth f)" -definition%important eval_fps :: "'a :: {banach, real_normed_div_algebra} fps \ 'a \ 'a" where +definition\<^marker>\tag important\ eval_fps :: "'a :: {banach, real_normed_div_algebra} fps \ 'a \ 'a" where "eval_fps f z = (\n. fps_nth f n * z ^ n)" -definition%important fps_expansion :: "(complex \ complex) \ complex \ complex fps" where +definition\<^marker>\tag important\ fps_expansion :: "(complex \ complex) \ complex \ complex fps" where "fps_expansion f z0 = Abs_fps (\n. (deriv ^^ n) f z0 / fact n)" lemma norm_summable_fps: @@ -213,7 +213,7 @@ qed -subsection%unimportant \Lower bounds on radius of convergence\ +subsection\<^marker>\tag unimportant\ \Lower bounds on radius of convergence\ 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. \ -definition%important +definition\<^marker>\tag important\ has_fps_expansion :: "('a :: {banach,real_normed_div_algebra} \ 'a) \ 'a fps \ bool" (infixl "has'_fps'_expansion" 60) where "(f has_fps_expansion F) \ diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Fashoda_Theorem.thy --- 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 \Bijections between intervals\ -definition%important interval_bij :: "'a \ 'a \ 'a \ 'a \ 'a \ 'a::euclidean_space" +definition\<^marker>\tag important\ interval_bij :: "'a \ 'a \ 'a \ 'a \ 'a \ 'a::euclidean_space" where "interval_bij = (\(a, b) (u, v) x. (\i\Basis. (u\i + (x\i - a\i) / (b\i - a\i) * (v\i - u\i)) *\<^sub>R i))" @@ -470,7 +470,7 @@ qed -subsection%unimportant \Some slightly ad hoc lemmas I use below\ +subsection\<^marker>\tag unimportant\ \Some slightly ad hoc lemmas I use below\ lemma segment_vertical: fixes a :: "real^2" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Finite_Cartesian_Product.thy --- 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 \Finite Cartesian products, with indexing and lambdas\ +subsection\<^marker>\tag unimportant\ \Finite Cartesian products, with indexing and lambdas\ typedef ('a, 'b) vec = "UNIV :: ('b::finite \ 'a) set" morphisms vec_nth vec_lambda .. @@ -165,7 +165,7 @@ by (simp add: image_comp o_def vec_nth_inverse) qed -subsection%unimportant \Group operations and class instances\ +subsection\<^marker>\tag unimportant\ \Group operations and class instances\ instantiation vec :: (zero, finite) zero begin @@ -230,7 +230,7 @@ by standard (simp_all add: vec_eq_iff) -subsection%unimportant\Basic componentwise operations on vectors\ +subsection\<^marker>\tag unimportant\\Basic componentwise operations on vectors\ instantiation vec :: (times, finite) times begin @@ -294,15 +294,15 @@ subsection \Real vector space\ -instantiation%unimportant vec :: (real_vector, finite) real_vector +instantiation\<^marker>\tag unimportant\ vec :: (real_vector, finite) real_vector begin -definition%important "scaleR \ (\ r x. (\ i. scaleR r (x$i)))" +definition\<^marker>\tag important\ "scaleR \ (\ r x. (\ 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>\tag unimportant\ by standard (simp_all add: vec_eq_iff scaleR_left_distrib scaleR_right_distrib) end @@ -310,15 +310,15 @@ subsection \Topological space\ -instantiation%unimportant vec :: (topological_space, finite) topological_space +instantiation\<^marker>\tag unimportant\ vec :: (topological_space, finite) topological_space begin -definition%important [code del]: +definition\<^marker>\tag important\ [code del]: "open (S :: ('a ^ 'b) set) \ (\x\S. \A. (\i. open (A i) \ x$i \ A i) \ (\y. (\i. y$i \ A i) \ y \ S))" -instance%unimportant proof +instance\<^marker>\tag unimportant\ proof show "open (UNIV :: ('a ^ 'b) set)" unfolding open_vec_def by auto next @@ -418,7 +418,7 @@ then show "\T. open T \ a \ T \ T \ (\x. x $ i) ` S" by - (rule exI) qed -instance%unimportant vec :: (perfect_space, finite) perfect_space +instance\<^marker>\tag unimportant\ vec :: (perfect_space, finite) perfect_space proof fix x :: "'a ^ 'b" show "\ open {x}" proof @@ -433,29 +433,29 @@ subsection \Metric space\ (* TODO: Product of uniform spaces and compatibility with metric_spaces! *) -instantiation%unimportant vec :: (metric_space, finite) dist +instantiation\<^marker>\tag unimportant\ vec :: (metric_space, finite) dist begin -definition%important +definition\<^marker>\tag important\ "dist x y = L2_set (\i. dist (x$i) (y$i)) UNIV" instance .. end -instantiation%unimportant vec :: (metric_space, finite) uniformity_dist +instantiation\<^marker>\tag unimportant\ vec :: (metric_space, finite) uniformity_dist begin -definition%important [code del]: +definition\<^marker>\tag important\ [code del]: "(uniformity :: (('a^'b::_) \ ('a^'b::_)) filter) = (INF e\{0 <..}. principal {(x, y). dist x y < e})" -instance%unimportant +instance\<^marker>\tag unimportant\ 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>\tag unimportant\ vec :: (metric_space, finite) metric_space begin proposition dist_vec_nth_le: "dist (x $ i) (y $ i) \ dist x y" @@ -557,7 +557,7 @@ then show "\M. \m\M. \n\M. dist (X m) (X n) < r" .. qed -instance%unimportant vec :: (complete_space, finite) complete_space +instance\<^marker>\tag unimportant\ vec :: (complete_space, finite) complete_space proof fix X :: "nat \ 'a ^ 'b" assume "Cauchy X" have "\i. (\n. X n $ i) \ lim (\n. X n $ i)" @@ -572,14 +572,14 @@ subsection \Normed vector space\ -instantiation%unimportant vec :: (real_normed_vector, finite) real_normed_vector +instantiation\<^marker>\tag unimportant\ vec :: (real_normed_vector, finite) real_normed_vector begin -definition%important "norm x = L2_set (\i. norm (x$i)) UNIV" +definition\<^marker>\tag important\ "norm x = L2_set (\i. norm (x$i)) UNIV" -definition%important "sgn (x::'a^'b) = scaleR (inverse (norm x)) x" +definition\<^marker>\tag important\ "sgn (x::'a^'b) = scaleR (inverse (norm x)) x" -instance%unimportant proof +instance\<^marker>\tag unimportant\ proof fix a :: real and x y :: "'a ^ 'b" show "norm x = 0 \ x = 0" unfolding norm_vec_def @@ -609,8 +609,8 @@ fixes x :: "'a::real_normed_vector^'n" assumes "\i. norm(x$i) \ norm(y$i)" shows "norm x \ 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: "\x$i\ \ norm x" apply (simp add: norm_vec_def) @@ -638,12 +638,12 @@ subsection \Inner product space\ -instantiation%unimportant vec :: (real_inner, finite) real_inner +instantiation\<^marker>\tag unimportant\ vec :: (real_inner, finite) real_inner begin -definition%important "inner x y = sum (\i. inner (x$i) (y$i)) UNIV" +definition\<^marker>\tag important\ "inner x y = sum (\i. inner (x$i) (y$i)) UNIV" -instance%unimportant proof +instance\<^marker>\tag unimportant\ 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 \Vectors pointing along a single axis.\ -definition%important "axis k x = (\ i. if i = k then x else 0)" +definition\<^marker>\tag important\ "axis k x = (\ 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>\tag unimportant\ vec :: (euclidean_space, finite) euclidean_space begin -definition%important "Basis = (\i. \u\Basis. {axis i u})" +definition\<^marker>\tag important\ "Basis = (\i. \u\Basis. {axis i u})" -instance%unimportant proof +instance\<^marker>\tag unimportant\ proof show "(Basis :: ('a ^ 'b) set) \ {}" 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 \A naive proof procedure to lift really trivial arithmetic stuff from the basis of the vector space\ +subsection\<^marker>\tag unimportant\ \A naive proof procedure to lift really trivial arithmetic stuff from the basis of the vector space\ lemma sum_cong_aux: "(\x. x \ A \ f x = g x) \ 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>\tag unimportant\ 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 \Some frequently useful arithmetic lemmas over vectors\ +subsection\<^marker>\tag unimportant\ \Some frequently useful arithmetic lemmas over vectors\ instance vec :: (semigroup_mult, finite) semigroup_mult by standard (vector mult.assoc) @@ -996,31 +996,31 @@ text\Matrix notation. NB: an MxN matrix is of type \<^typ>\'a^'n^'m\, not \<^typ>\'a^'m^'n\\ -definition%important map_matrix::"('a \ 'b) \ (('a, 'i::finite)vec, 'j::finite) vec \ (('b, 'i)vec, 'j) vec" where +definition\<^marker>\tag important\ map_matrix::"('a \ 'b) \ (('a, 'i::finite)vec, 'j::finite) vec \ (('b, 'i)vec, 'j) vec" where "map_matrix f x = (\ 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 \ 'a ^'p^'n \ 'a ^ 'p ^'m" +definition\<^marker>\tag important\ matrix_matrix_mult :: "('a::semiring_1) ^'n^'m \ 'a ^'p^'n \ 'a ^ 'p ^'m" (infixl "**" 70) where "m ** m' == (\ i j. sum (\k. ((m$i)$k) * ((m'$k)$j)) (UNIV :: 'n set)) ::'a ^ 'p ^'m" -definition%important matrix_vector_mult :: "('a::semiring_1) ^'n^'m \ 'a ^'n \ 'a ^ 'm" +definition\<^marker>\tag important\ matrix_vector_mult :: "('a::semiring_1) ^'n^'m \ 'a ^'n \ 'a ^ 'm" (infixl "*v" 70) where "m *v x \ (\ i. sum (\j. ((m$i)$j) * (x$j)) (UNIV ::'n set)) :: 'a^'m" -definition%important vector_matrix_mult :: "'a ^ 'm \ ('a::semiring_1) ^'n^'m \ 'a ^'n " +definition\<^marker>\tag important\ vector_matrix_mult :: "'a ^ 'm \ ('a::semiring_1) ^'n^'m \ 'a ^'n " (infixl "v*" 70) where "v v* m == (\ j. sum (\i. ((m$i)$j) * (v$i)) (UNIV :: 'm set)) :: 'a^'n" -definition%unimportant "(mat::'a::zero => 'a ^'n^'n) k = (\ i j. if i = j then k else 0)" -definition%unimportant transpose where +definition\<^marker>\tag unimportant\ "(mat::'a::zero => 'a ^'n^'n) k = (\ i j. if i = j then k else 0)" +definition\<^marker>\tag unimportant\ transpose where "(transpose::'a^'n^'m \ 'a^'m^'n) A = (\ i j. ((A$j)$i))" -definition%unimportant "(row::'m => 'a ^'n^'m \ 'a ^'n) i A = (\ j. ((A$i)$j))" -definition%unimportant "(column::'n =>'a^'n^'m =>'a^'m) j A = (\ i. ((A$i)$j))" -definition%unimportant "rows(A::'a^'n^'m) = { row i A | i. i \ (UNIV :: 'm set)}" -definition%unimportant "columns(A::'a^'n^'m) = { column i A | i. i \ (UNIV :: 'n set)}" +definition\<^marker>\tag unimportant\ "(row::'m => 'a ^'n^'m \ 'a ^'n) i A = (\ j. ((A$i)$j))" +definition\<^marker>\tag unimportant\ "(column::'n =>'a^'n^'m =>'a^'m) j A = (\ i. ((A$i)$j))" +definition\<^marker>\tag unimportant\ "rows(A::'a^'n^'m) = { row i A | i. i \ (UNIV :: 'm set)}" +definition\<^marker>\tag unimportant\ "columns(A::'a^'n^'m) = { column i A | i. i \ (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\Correspondence between matrices and linear operators.\ -definition%important matrix :: "('a::{plus,times, one, zero}^'m \ 'a ^ 'n) \ 'a^'m^'n" +definition\<^marker>\tag important\ matrix :: "('a::{plus,times, one, zero}^'m \ 'a ^ 'n) \ 'a^'m^'n" where "matrix f = (\ i j. (f(axis j 1))$i)" lemma matrix_id_mat_1: "matrix id = mat 1" @@ -1204,10 +1204,10 @@ subsection\Inverse matrices (not necessarily square)\ -definition%important +definition\<^marker>\tag important\ "invertible(A::'a::semiring_1^'n^'m) \ (\A'::'a^'m^'n. A ** A' = mat 1 \ A' ** A = mat 1)" -definition%important +definition\<^marker>\tag important\ "matrix_inv(A:: 'a::semiring_1^'n^'m) = (SOME A'::'a^'m^'n. A ** A' = mat 1 \ A' ** A = mat 1)" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Finite_Product_Measure.thy --- 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: "(\(i, j). c) = (\_. c)" by auto -subsection%unimportant \More about Function restricted by \<^const>\extensional\\ +subsection\<^marker>\tag unimportant\ \More about Function restricted by \<^const>\extensional\\ definition "merge I J = (\(x, y) i. if i \ I then x i else if i \ J then y i else undefined)" @@ -111,12 +111,12 @@ subsection \Finite product spaces\ -definition%important prod_emb where +definition\<^marker>\tag important\ prod_emb where "prod_emb I M K X = (\x. restrict x K) -` X \ (\\<^sub>E i\I. space (M i))" lemma prod_emb_iff: "f \ prod_emb I M K X \ f \ extensional I \ (restrict f K \ X) \ (\i\I. f i \ 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 \ G \ prod_emb A M B F \ prod_emb A M B G" by (auto simp: prod_emb_def) -definition%important PiM :: "'i set \ ('i \ 'a measure) \ ('i \ 'a) measure" where +definition\<^marker>\tag important\ PiM :: "'i set \ ('i \ 'a measure) \ ('i \ 'a) measure" where "PiM I M = extend_measure (\\<^sub>E i\I. space (M i)) {(J, X). (J \ {} \ I = {}) \ finite J \ J \ I \ X \ (\ j\J. sets (M j))} (\(J, X). prod_emb I M J (\\<^sub>E j\J. X j)) (\(J, X). \j\J \ {i\I. emeasure (M i) (space (M i)) \ 1}. if j \ J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))" -definition%important prod_algebra :: "'i set \ ('i \ 'a measure) \ ('i \ 'a) set set" where +definition\<^marker>\tag important\ prod_algebra :: "'i set \ ('i \ 'a measure) \ ('i \ 'a) set set" where "prod_algebra I M = (\(J, X). prod_emb I M J (\\<^sub>E j\J. X j)) ` {(J, X). (J \ {} \ I = {}) \ finite J \ J \ I \ X \ (\ j\J. sets (M j))}" @@ -825,14 +825,14 @@ by (auto intro!: eq) qed (auto simp: sets_PiM) -locale%unimportant product_sigma_finite = +locale\<^marker>\tag unimportant\ product_sigma_finite = fixes M :: "'i \ 'a measure" assumes sigma_finite_measures: "\i. sigma_finite_measure (M i)" -sublocale%unimportant product_sigma_finite \ M?: sigma_finite_measure "M i" for i +sublocale\<^marker>\tag unimportant\ product_sigma_finite \ 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 \ 'a measure" + +locale\<^marker>\tag unimportant\ finite_product_sigma_finite = product_sigma_finite M for M :: "'i \ 'a measure" + fixes I :: "'i set" assumes finite_index: "finite I" @@ -991,7 +991,7 @@ "0 \ f (\k. undefined) \ integral\<^sup>N (Pi\<^sub>M {} M) f = f (\k. undefined)" by (simp add: PiM_empty nn_integral_count_space_finite max.absorb2) -lemma%important (in product_sigma_finite) distr_merge: +lemma\<^marker>\tag important\ (in product_sigma_finite) distr_merge: assumes IJ[simp]: "I \ J = {}" and fin: "finite I" "finite J" shows "distr (Pi\<^sub>M I M \\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \ J) M) (merge I J) = Pi\<^sub>M (I \ J) M" (is "?D = ?P") diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Function_Topology.thy --- 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. \ -definition%important product_topology::"('i \ ('a topology)) \ ('i set) \ (('i \ 'a) topology)" +definition\<^marker>\tag important\ product_topology::"('i \ ('a topology)) \ ('i set) \ (('i \ 'a) topology)" where "product_topology T I = topology_generated_by {(\\<^sub>E i\I. X i) |X. (\i. openin (T i) (X i)) \ finite {i. X i \ topspace (T i)}}" @@ -548,12 +548,12 @@ qed -subsubsection%important \Powers of a single topological space as a topological space, using type classes\ +subsubsection\<^marker>\tag important\ \Powers of a single topological space as a topological space, using type classes\ instantiation "fun" :: (type, topological_space) topological_space begin -definition%important open_fun_def: +definition\<^marker>\tag important\ open_fun_def: "open U = openin (product_topology (\i. euclidean) UNIV) U" instance proof @@ -653,7 +653,7 @@ by (intro continuous_intros) auto qed -subsubsection%important \Topological countability for product spaces\ +subsubsection\<^marker>\tag important\ \Topological countability for product spaces\ text \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>\tag important\ dist_fun_def: "dist x y = (\n. (1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1)" -definition%important uniformity_fun_def: +definition\<^marker>\tag important\ uniformity_fun_def: "(uniformity::(('a \ 'b) \ ('a \ 'b)) filter) = (INF e\{0<..}. principal {(x, y). dist (x::('a\'b)) y < e})" text \Except for the first one, the auxiliary lemmas below are only useful when proving the diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Further_Topology.thy --- 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 \ {}" shows "DIM('a) \ 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 \ 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 \ 'b::euclidean_space" @@ -2364,8 +2364,8 @@ fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "DIM('b) \ 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 \ 'b::euclidean_space" @@ -3375,7 +3375,7 @@ corollary covering_space_square_punctured_plane: "covering_space (- {0}) (\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\This doesn't have a standard name. Kuratowski uses ``contractible with respect to \[S\<^sup>1]\'' while Whyburn uses ``property b''. It's closely related to unicoherence.\ -definition%important Borsukian where +definition\<^marker>\tag important\ Borsukian where "Borsukian S \ \f. continuous_on S f \ f ` S \ (- {0::complex}) \ (\a. homotopic_with_canon (\h. True) S (- {0}) f (\x. a))" @@ -4815,7 +4815,7 @@ subsection\Unicoherence (closed)\ -definition%important unicoherent where +definition\<^marker>\tag important\ unicoherent where "unicoherent U \ \S T. connected S \ connected T \ S \ T = U \ closedin (top_of_set U) S \ 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\If the type class constraint can be relaxed, I don't know how!\ 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: diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Gamma_Function.thy --- 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. \ -definition%important Gamma_series :: "('a :: {banach,real_normed_field}) \ nat \ 'a" where +definition\<^marker>\tag important\ Gamma_series :: "('a :: {banach,real_normed_field}) \ nat \ '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}) \ nat \ 'a" where @@ -317,10 +317,10 @@ function to the inverse of the Gamma function, and from that to the Gamma function itself. \ -definition%important ln_Gamma_series :: "('a :: {banach,real_normed_field,ln}) \ nat \ 'a" where +definition\<^marker>\tag important\ ln_Gamma_series :: "('a :: {banach,real_normed_field,ln}) \ nat \ 'a" where "ln_Gamma_series z n = z * ln (of_nat n) - ln z - (\k=1..n. ln (z / of_nat k + 1))" -definition%unimportant ln_Gamma_series' :: "('a :: {banach,real_normed_field,ln}) \ nat \ 'a" where +definition\<^marker>\tag unimportant\ ln_Gamma_series' :: "('a :: {banach,real_normed_field,ln}) \ nat \ 'a" where "ln_Gamma_series' z n = - euler_mascheroni*z - ln z + (\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 \ ('a :: {real_normed_field,banach}) \ 'a" where +definition\<^marker>\tag important\ Polygamma :: "nat \ ('a :: {real_normed_field,banach}) \ 'a" where "Polygamma n z = (if n = 0 then (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) - euler_mascheroni else (-1)^Suc n * fact n * (\k. inverse ((z + of_nat k)^Suc n)))" -abbreviation%important Digamma :: "('a :: {real_normed_field,banach}) \ 'a" where +abbreviation\<^marker>\tag important\ Digamma :: "('a :: {real_normed_field,banach}) \ 'a" where "Digamma \ Polygamma 0" lemma Digamma_def: @@ -1021,7 +1021,7 @@ the reals and for the complex numbers with a minimal amount of proof duplication. \ -class%unimportant Gamma = real_normed_field + complete_space + +class\<^marker>\tag unimportant\ Gamma = real_normed_field + complete_space + fixes rGamma :: "'a \ 'a" assumes rGamma_eq_zero_iff_aux: "rGamma z = 0 \ (\n. z = - of_nat n)" assumes differentiable_rGamma_aux1: @@ -1259,12 +1259,12 @@ "isCont f z \ f z \ \\<^sub>\\<^sub>0 \ isCont (\x. Gamma (f x)) z" by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_Gamma]]) -subsection%unimportant \The complex Gamma function\ - -instantiation%unimportant complex :: Gamma +subsection\<^marker>\tag unimportant\ \The complex Gamma function\ + +instantiation\<^marker>\tag unimportant\ complex :: Gamma begin -definition%unimportant rGamma_complex :: "complex \ complex" where +definition\<^marker>\tag unimportant\ rGamma_complex :: "complex \ 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>\tag unimportant\ begin (* TODO: duplication *) @@ -1526,7 +1526,7 @@ -subsection%unimportant \The real Gamma function\ +subsection\<^marker>\tag unimportant\ \The real Gamma function\ lemma rGamma_series_real: "eventually (\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>\tag unimportant\ 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). \ -context%unimportant +context\<^marker>\tag unimportant\ fixes G :: "real \ real" assumes G_1: "G 1 = 1" assumes G_plus1: "x > 0 \ G (x + 1) = x * G x" @@ -2389,7 +2389,7 @@ end -subsection%unimportant \Limits and residues\ +subsection\<^marker>\tag unimportant\ \Limits and residues\ text \ The inverse of the Gamma function has simple zeros: @@ -2537,7 +2537,7 @@ "Gamma_series_Weierstrass z n = exp (-euler_mascheroni * z) / z * (\k=1..n. exp (z / of_nat k) / (1 + z / of_nat k))" -definition%unimportant +definition\<^marker>\tag unimportant\ rGamma_series_Weierstrass :: "'a :: {banach,real_normed_field} \ nat \ 'a" where "rGamma_series_Weierstrass z n = exp (euler_mascheroni * z) * z * (\k=1..n. (1 + z / of_nat k) * exp (-z / of_nat k))" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Great_Picard.thy --- 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\Montel's theorem\ +subsubsection\<^marker>\tag important\\Montel's theorem\ text\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 \ M" and holf: "f holomorphic_on (M - {z})" and non: "\l. \ (f \ l) (at z)" "\l. \ ((inverse \ f) \ l) (at z)" obtains a where "- {a} \ 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: diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Harmonic_Numbers.thy --- 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 \The Harmonic numbers\ -definition%important harm :: "nat \ 'a :: real_normed_field" where +definition\<^marker>\tag important\ harm :: "nat \ 'a :: real_normed_field" where "harm n = (\k=1..n. inverse (of_nat k))" lemma harm_altdef: "harm n = (\k\tag important\ euler_mascheroni_LIMSEQ: "(\n. harm n - ln (of_nat n) :: real) \ euler_mascheroni" unfolding euler_mascheroni_def by (simp add: convergent_LIMSEQ_iff [symmetric] euler_mascheroni_convergent) @@ -250,7 +250,7 @@ qed -subsection%unimportant \Bounds on the Euler-Mascheroni constant\ +subsection\<^marker>\tag unimportant\ \Bounds on the Euler-Mascheroni constant\ (* TODO: Move? *) lemma ln_inverse_approx_le: diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Homeomorphism.thy --- 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 \ 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\Covering spaces and lifting results for them\ -definition%important covering_space +definition\<^marker>\tag important\ covering_space :: "'a::topological_space set \ ('a \ 'b) \ 'b::topological_space set \ bool" where "covering_space c p S \ @@ -2130,8 +2130,8 @@ and "path h2" and pih2: "path_image h2 \ C" and ph2: "\t. t \ {0..1} \ 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: diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Homotopy.thy --- 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>\tag important\ homotopic_with where "homotopic_with P X Y f g \ (\h. continuous_map (prod_topology (subtopology euclideanreal {0..1}) X) Y h \ @@ -50,10 +50,10 @@ apply (simp add: t) done -subsection%unimportant\Trivial properties\ +subsection\<^marker>\tag unimportant\\Trivial properties\ text \We often want to just localize the ending function equality or whatever.\ -text%important \%whitespace\ +text\<^marker>\tag important\ \%whitespace\ proposition homotopic_with: assumes "\h k. (\x. x \ topspace X \ h x = k x) \ (P h \ P k)" shows "homotopic_with P X Y p q \ @@ -516,7 +516,7 @@ subsection\Homotopy of paths, maintaining the same endpoints\ -definition%important homotopic_paths :: "['a set, real \ 'a, real \ 'a::topological_space] \ bool" +definition\<^marker>\tag important\ homotopic_paths :: "['a set, real \ 'a, real \ 'a::topological_space] \ bool" where "homotopic_paths s p q \ homotopic_with_canon (\r. pathstart r = pathstart p \ pathfinish r = pathfinish p) {0..1} s p q" @@ -673,7 +673,7 @@ subsection\Group properties for homotopy of paths\ -text%important\So taking equivalence classes under homotopy would give the fundamental group\ +text\<^marker>\tag important\\So taking equivalence classes under homotopy would give the fundamental group\ proposition homotopic_paths_rid: "\path p; path_image p \ s\ \ homotopic_paths s (p +++ linepath (pathfinish p) (pathfinish p)) p" @@ -736,7 +736,7 @@ subsection\Homotopy of loops without requiring preservation of endpoints\ -definition%important homotopic_loops :: "'a::topological_space set \ (real \ 'a) \ (real \ 'a) \ bool" where +definition\<^marker>\tag important\ homotopic_loops :: "'a::topological_space set \ (real \ 'a) \ (real \ 'a) \ bool" where "homotopic_loops s p q \ homotopic_with_canon (\r. pathfinish r = pathstart r) {0..1} s p q" @@ -971,7 +971,7 @@ qed -subsection%unimportant\Homotopy of "nearby" function, paths and loops\ +subsection\<^marker>\tag unimportant\\Homotopy of "nearby" function, paths and loops\ lemma homotopic_with_linear: fixes f g :: "_ \ 'b::real_normed_vector" @@ -1173,9 +1173,9 @@ subsection\Simply connected sets\ -text%important\defined as "all loops are homotopic (as loops)\ - -definition%important simply_connected where +text\<^marker>\tag important\\defined as "all loops are homotopic (as loops)\ + +definition\<^marker>\tag important\ simply_connected where "simply_connected S \ \p q. path p \ pathfinish p = pathstart p \ path_image p \ S \ path q \ pathfinish q = pathstart q \ path_image q \ S @@ -1357,7 +1357,7 @@ subsection\Contractible sets\ -definition%important contractible where +definition\<^marker>\tag important\ contractible where "contractible S \ \a. homotopic_with_canon (\x. True) S S id (\x. a)" proposition contractible_imp_simply_connected: @@ -1598,7 +1598,7 @@ subsection\Local versions of topological properties in general\ -definition%important locally :: "('a::topological_space set \ bool) \ 'a set \ bool" +definition\<^marker>\tag important\ locally :: "('a::topological_space set \ bool) \ 'a set \ bool" where "locally P S \ \w x. openin (top_of_set S) w \ x \ w @@ -2975,7 +2975,7 @@ by metis qed -subsection%unimportant\Components, continuity, openin, closedin\ +subsection\<^marker>\tag unimportant\\Components, continuity, openin, closedin\ lemma continuous_on_components_gen: fixes f :: "'a::topological_space \ 'b::topological_space" @@ -3308,7 +3308,7 @@ subsection\Retracts, in a general sense, preserve (co)homotopic triviality)\ -locale%important Retracts = +locale\<^marker>\tag important\ Retracts = fixes s h t k assumes conth: "continuous_on s h" and imh: "h ` s = t" @@ -3474,7 +3474,7 @@ subsection\Homotopy equivalence of topological spaces.\ -definition%important homotopy_equivalent_space +definition\<^marker>\tag important\ homotopy_equivalent_space (infix "homotopy'_equivalent'_space" 50) where "X homotopy_equivalent_space Y \ (\f g. continuous_map X Y f \ @@ -3897,7 +3897,7 @@ qed -abbreviation%important homotopy_eqv :: "'a::topological_space set \ 'b::topological_space set \ bool" +abbreviation\<^marker>\tag important\ homotopy_eqv :: "'a::topological_space set \ 'b::topological_space set \ bool" (infix "homotopy'_eqv" 50) where "S homotopy_eqv T \ top_of_set S homotopy_equivalent_space top_of_set T" @@ -4144,7 +4144,7 @@ by (metis homeomorphic_contractible_eq) -subsection%unimportant\Misc other results\ +subsection\<^marker>\tag unimportant\\Misc other results\ lemma bounded_connected_Compl_real: fixes S :: "real set" @@ -4194,7 +4194,7 @@ qed -subsection%unimportant\Some Uncountable Sets\ +subsection\<^marker>\tag unimportant\\Some Uncountable Sets\ 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\ Some simple positive connection theorems\ +subsection\<^marker>\tag unimportant\\ Some simple positive connection theorems\ proposition path_connected_convex_diff_countable: fixes U :: "'a::euclidean_space set" @@ -4598,9 +4598,9 @@ -subsection%unimportant \Self-homeomorphisms shuffling points about\ - -subsubsection%unimportant\The theorem \homeomorphism_moving_points_exists\\ +subsection\<^marker>\tag unimportant\ \Self-homeomorphisms shuffling points about\ + +subsubsection\<^marker>\tag unimportant\\The theorem \homeomorphism_moving_points_exists\\ lemma homeomorphism_moving_point_1: fixes a :: "'a::euclidean_space" @@ -4730,7 +4730,7 @@ done qed -corollary%unimportant homeomorphism_moving_point_2: +corollary\<^marker>\tag unimportant\ homeomorphism_moving_point_2: fixes a :: "'a::euclidean_space" assumes "affine T" "a \ T" and u: "u \ ball a r \ T" and v: "v \ ball a r \ T" obtains f g where "homeomorphism (cball a r \ T) (cball a r \ T) f g" @@ -4758,7 +4758,7 @@ qed -corollary%unimportant homeomorphism_moving_point_3: +corollary\<^marker>\tag unimportant\ homeomorphism_moving_point_3: fixes a :: "'a::euclidean_space" assumes "affine T" "a \ T" and ST: "ball a r \ T \ S" "S \ T" and u: "u \ ball a r \ T" and v: "v \ ball a r \ T" @@ -4836,7 +4836,7 @@ qed -proposition%unimportant homeomorphism_moving_point: +proposition\<^marker>\tag unimportant\ homeomorphism_moving_point: fixes a :: "'a::euclidean_space" assumes ope: "openin (top_of_set (affine hull S)) S" and "S \ T" @@ -4992,7 +4992,7 @@ qed qed -proposition%unimportant homeomorphism_moving_points_exists: +proposition\<^marker>\tag unimportant\ homeomorphism_moving_points_exists: fixes S :: "'a::euclidean_space set" assumes 2: "2 \ DIM('a)" "open S" "connected S" "S \ T" "finite K" and KS: "\i. i \ K \ x i \ S \ y i \ S" @@ -5021,7 +5021,7 @@ using homeomorphism_moving_points_exists_gen [OF \finite K\ KS pw _ ope S] that by fastforce qed -subsubsection%unimportant\The theorem \homeomorphism_grouping_points_exists\\ +subsubsection\<^marker>\tag unimportant\\The theorem \homeomorphism_grouping_points_exists\\ 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>\tag unimportant\ homeomorphism_grouping_points_exists: fixes S :: "'a::euclidean_space set" assumes "open U" "open S" "connected S" "U \ {}" "finite K" "K \ S" "U \ S" "S \ T" obtains f g where "homeomorphism T T f g" "{x. (\ (f x = x \ g x = x))} \ S" @@ -5337,7 +5337,7 @@ qed -proposition%unimportant homeomorphism_grouping_points_exists_gen: +proposition\<^marker>\tag unimportant\ 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" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Improper_Integral.thy --- 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\The definition here only really makes sense for an elementary set. We just use compact intervals in applications below.\ -definition%important equiintegrable_on (infixr "equiintegrable'_on" 46) +definition\<^marker>\tag important\ equiintegrable_on (infixr "equiintegrable'_on" 46) where "F equiintegrable_on I \ (\f \ F. f integrable_on I) \ (\e > 0. \\. gauge \ \ diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Infinite_Products.thy --- 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 \Preliminaries\ +subsection\<^marker>\tag unimportant\ \Preliminaries\ lemma sum_le_prod: fixes f :: "'a \ 'b :: linordered_semidom" @@ -54,19 +54,19 @@ subsection\Definitions and basic properties\ -definition%important raw_has_prod :: "[nat \ 'a::{t2_space, comm_semiring_1}, nat, 'a] \ bool" +definition\<^marker>\tag important\ raw_has_prod :: "[nat \ 'a::{t2_space, comm_semiring_1}, nat, 'a] \ bool" where "raw_has_prod f M p \ (\n. \i\n. f (i+M)) \ p \ p \ 0" text\The nonzero and zero cases, as in \emph{Complex Analysis} by Joseph Bak and Donald J.Newman, page 241\ -text%important \%whitespace\ -definition%important +text\<^marker>\tag important\ \%whitespace\ +definition\<^marker>\tag important\ has_prod :: "(nat \ 'a::{t2_space, comm_semiring_1}) \ 'a \ bool" (infixr "has'_prod" 80) where "f has_prod p \ raw_has_prod f 0 p \ (\i q. p = 0 \ f i = 0 \ raw_has_prod f (Suc i) q)" -definition%important convergent_prod :: "(nat \ 'a :: {t2_space,comm_semiring_1}) \ bool" where +definition\<^marker>\tag important\ convergent_prod :: "(nat \ 'a :: {t2_space,comm_semiring_1}) \ bool" where "convergent_prod f \ \M p. raw_has_prod f M p" -definition%important prodinf :: "(nat \ 'a::{t2_space, comm_semiring_1}) \ 'a" +definition\<^marker>\tag important\ prodinf :: "(nat \ 'a::{t2_space, comm_semiring_1}) \ 'a" (binder "\" 10) where "prodinf f = (THE p. f has_prod p)" @@ -148,7 +148,7 @@ subsection\Absolutely convergent products\ -definition%important abs_convergent_prod :: "(nat \ _) \ bool" where +definition\<^marker>\tag important\ abs_convergent_prod :: "(nat \ _) \ bool" where "abs_convergent_prod f \ convergent_prod (\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>\tag important\ convergent_prod_iff_convergent: fixes f :: "nat \ 'a :: {topological_semigroup_mult,t2_space,idom}" assumes "\i. f i \ 0" shows "convergent_prod f \ convergent (\n. \i\n. f i) \ lim (\n. \i\n. f i) \ 0" @@ -398,7 +398,7 @@ thus ?thesis by eventually_elim auto qed -subsection%unimportant \Ignoring initial segments\ +subsection\<^marker>\tag unimportant\ \Ignoring initial segments\ lemma convergent_prod_offset: assumes "convergent_prod (\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>\tag unimportant\ convergent_prod_ignore_initial_segment: fixes f :: "nat \ 'a :: real_normed_field" assumes "convergent_prod f" shows "convergent_prod (\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>\tag unimportant\ convergent_prod_ignore_nonzero_segment: fixes f :: "nat \ 'a :: real_normed_field" assumes f: "convergent_prod f" and nz: "\i. i \ M \ f i \ 0" shows "\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>\tag unimportant\ abs_convergent_prod_ignore_initial_segment: assumes "abs_convergent_prod f" shows "abs_convergent_prod (\n. f (n + m))" using assms unfolding abs_convergent_prod_def @@ -761,7 +761,7 @@ qed qed -corollary%unimportant has_prod_0: +corollary\<^marker>\tag unimportant\ has_prod_0: fixes f :: "nat \ 'a::{semidom,t2_space}" assumes "\n. f n = 1" shows "f has_prod 1" @@ -874,7 +874,7 @@ end -subsection%unimportant \Infinite products on ordered topological monoids\ +subsection\<^marker>\tag unimportant\ \Infinite products on ordered topological monoids\ lemma LIMSEQ_prod_0: fixes f :: "nat \ 'a::{semidom,topological_space}" @@ -1075,7 +1075,7 @@ using convergent_prod_def raw_has_prodI_bounded [OF assms] by blast -subsection%unimportant \Infinite products on topological spaces\ +subsection\<^marker>\tag unimportant\ \Infinite products on topological spaces\ context fixes f g :: "nat \ 'a::{t2_space,topological_semigroup_mult,idom}" @@ -1176,7 +1176,7 @@ end -subsection%unimportant \Infinite summability on real normed fields\ +subsection\<^marker>\tag unimportant\ \Infinite summability on real normed fields\ context fixes f :: "nat \ 'a::real_normed_field" @@ -1332,7 +1332,7 @@ by (simp add: ac_simps) qed -corollary%unimportant has_prod_iff_shift': +corollary\<^marker>\tag unimportant\ has_prod_iff_shift': assumes "\i. i < n \ f i \ 0" shows "(\i. f (i + n)) has_prod (a / (\i 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 \Embeddings from the reals into some complete real normed field\ +subsection\<^marker>\tag unimportant\ \Embeddings from the reals into some complete real normed field\ lemma tendsto_eq_of_real_lim: assumes "(\n. of_real (f n) :: 'a::{complete_space,real_normed_field}) \ q" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Infinite_Set_Sum.thy --- 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>\tag important\ abs_summable_on :: "('a \ 'b :: {banach, second_countable_topology}) \ 'a set \ bool" (infix "abs'_summable'_on" 50) where "f abs_summable_on A \ integrable (count_space A) f" -definition%important infsetsum :: +definition\<^marker>\tag important\ infsetsum :: "('a \ 'b :: {banach, second_countable_topology}) \ 'a set \ 'b" where "infsetsum f A = lebesgue_integral (count_space A) f" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Inner_Product.thy --- 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 \Gradient derivative\ -definition%important +definition\<^marker>\tag important\ gderiv :: "['a::real_inner \ real, 'a, 'a] \ bool" ("(GDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Integral_Test.thy --- 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 @@ \ (* TODO: continuous_in \ integrable_on *) -locale%important antimono_fun_sum_integral_diff = +locale\<^marker>\tag important\ antimono_fun_sum_integral_diff = fixes f :: "real \ real" assumes dec: "\x y. x \ 0 \ x \ y \ f x \ f y" assumes nonneg: "\x. x \ 0 \ f x \ 0" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Interval_Integral.thy --- 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 \ ereal \ ereal \ (real \ 'a) \ 'a::{banach, second_countable_topology}" where +definition\<^marker>\tag important\ interval_lebesgue_integral :: "real measure \ ereal \ ereal \ (real \ 'a) \ 'a::{banach, second_countable_topology}" where "interval_lebesgue_integral M a b f = (if a \ 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 (\x. f)" -definition%important interval_lebesgue_integrable :: "real measure \ ereal \ ereal \ (real \ 'a::{banach, second_countable_topology}) \ bool" where +definition\<^marker>\tag important\ interval_lebesgue_integrable :: "real measure \ ereal \ ereal \ (real \ 'a::{banach, second_countable_topology}) \ bool" where "interval_lebesgue_integrable M a b f = (if a \ 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 \ f) = (set_integrable M {a<..} f)" - unfolding%unimportant interval_lebesgue_integrable_def by auto + unfolding interval_lebesgue_integrable_def by auto subsection\Basic properties of integration over an interval wrt lebesgue measure\ @@ -1083,8 +1083,8 @@ fixes f :: "real \ 'a :: {banach, second_countable_topology}" shows "interval_lebesgue_integrable lborel a b f \ a \ b \ norm (LBINT t=a..b. f t) \ LBINT t=a..b. norm (f t)" - using%unimportant integral_norm_bound[of lborel "\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 "\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 \ diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Jordan_Curve.thy --- 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>\tag unimportant\ Jordan_curve: fixes c :: "real \ 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>\tag unimportant\ Jordan_disconnected: fixes c :: "real \ complex" assumes "simple_path c" "pathfinish c = pathstart c" shows "\ connected(- path_image c)" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/L2_Norm.thy --- 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 \L2 Norm\ -definition%important L2_set :: "('a \ real) \ 'a set \ real" where +definition\<^marker>\tag important\ L2_set :: "('a \ real) \ 'a set \ real" where "L2_set f A = sqrt (\i\A. (f i)\<^sup>2)" lemma L2_set_cong: diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Lebesgue_Measure.thy --- 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: \ -definition%important interval_measure :: "(real \ real) \ real measure" where +definition\<^marker>\tag important\ interval_measure :: "(real \ real) \ real measure" where "interval_measure F = extend_measure UNIV {(a, b). a \ b} (\(a, b). {a<..b}) (\(a, b). ennreal (F b - F a))" -lemma%important emeasure_interval_measure_Ioc: +lemma\<^marker>\tag important\ emeasure_interval_measure_Ioc: assumes "a \ b" assumes mono_F: "\x y. x \ y \ F x \ F y" assumes right_cont_F : "\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 \a \ b\]) +proof (rule extend_measure_caratheodory_pair[OF interval_measure_def \a \ b\]) show "semiring_of_sets UNIV {{a<..b} |a b :: real. a \ b}" proof (unfold_locales, safe) fix a b c d :: real assume *: "a \ b" "c \ d" @@ -319,7 +319,7 @@ emeasure (interval_measure F) {a <.. b} = (if a \ 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>\tag important\ 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>\tag important\ sigma_finite_interval_measure: assumes mono_F: "\x y. x \ y \ F x \ F y" assumes right_cont_F : "\a. continuous (at_right a) F" shows "sigma_finite_measure (interval_measure F)" @@ -382,13 +382,13 @@ subsection \Lebesgue-Borel measure\ -definition%important lborel :: "('a :: euclidean_space) measure" where +definition\<^marker>\tag important\ lborel :: "('a :: euclidean_space) measure" where "lborel = distr (\\<^sub>M b\Basis. interval_measure (\x. x)) borel (\f. \b\Basis. f b *\<^sub>R b)" -abbreviation%important lebesgue :: "'a::euclidean_space measure" +abbreviation\<^marker>\tag important\ lebesgue :: "'a::euclidean_space measure" where "lebesgue \ completion lborel" -abbreviation%important lebesgue_on :: "'a set \ 'a::euclidean_space measure" +abbreviation\<^marker>\tag important\ lebesgue_on :: "'a set \ 'a::euclidean_space measure" where "lebesgue_on \ \ restrict_space (completion lborel) \" lemma @@ -409,7 +409,7 @@ shows "f \ measurable (lebesgue_on S) M \ g \ measurable (lebesgue_on S) M" by (metis (mono_tags, lifting) IntD1 assms measurable_cong_simp space_restrict_space) -text%unimportant \Measurability of continuous functions\ +text\<^marker>\tag unimportant\ \Measurability of continuous functions\ lemma continuous_imp_measurable_on_sets_lebesgue: assumes f: "continuous_on S f" and S: "S \ sets lebesgue" @@ -460,10 +460,10 @@ lemma emeasure_lborel_Icc_eq: "emeasure lborel {l .. u} = ennreal (if l \ u then u - l else 0)" by simp -lemma%important emeasure_lborel_cbox[simp]: +lemma\<^marker>\tag important\ emeasure_lborel_cbox[simp]: assumes [simp]: "\b. b \ Basis \ l \ b \ u \ b" shows "emeasure lborel (cbox l u) = (\b\Basis. (u - l) \ b)" -proof%unimportant - +proof - have "(\x. \b\Basis. indicator {l\b .. u\b} (x \ b) :: ennreal) = indicator (cbox l u)" by (auto simp: fun_eq_iff cbox_def split: split_indicator) then have "emeasure lborel (cbox l u) = (\\<^sup>+x. (\b\Basis. indicator {l\b .. u\b} (x \ b)) \lborel)" @@ -654,12 +654,12 @@ subsection \Affine transformation on the Lebesgue-Borel\ -lemma%important lborel_eqI: +lemma\<^marker>\tag important\ lborel_eqI: fixes M :: "'a::euclidean_space measure" assumes emeasure_eq: "\l u. (\b. b \ Basis \ l \ b \ u \ b) \ emeasure M (box l u) = (\b\Basis. (u - l) \ 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 (\(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>\tag important\ lborel_affine_euclidean: fixes c :: "'a::euclidean_space \ real" and t defines "T x \ t + (\j\Basis. (c j * (x \ j)) *\<^sub>R j)" assumes c: "\j. j \ Basis \ c j \ 0" shows "lborel = density (distr lborel borel T) (\_. (\j\Basis. \c j\))" (is "_ = ?D") -proof%unimportant (rule lborel_eqI) +proof (rule lborel_eqI) let ?B = "Basis :: 'a set" fix l u assume le: "\b. b \ ?B \ l \ b \ u \ b" have [measurable]: "T \ borel \\<^sub>M borel" @@ -736,10 +736,10 @@ lborel_integrable_real_affine[of "\x. f (t + c * x)" "1/c" "-t/c"] by (auto simp add: field_simps) -lemma%important lborel_integral_real_affine: +lemma\<^marker>\tag important\ lborel_integral_real_affine: fixes f :: "real \ 'a :: {banach, second_countable_topology}" and c :: real assumes c: "c \ 0" shows "(\x. f x \ lborel) = \c\ *\<^sub>R (\x. f (t + c * x) \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>\tag important\ lborel_prod: "lborel \\<^sub>M lborel = (lborel :: ('a::euclidean_space \ '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: "\a b. (a, b) \ Basis \ (la, lb) \ (a, b) \ (ua, ub) \ (a, b)" have [simp]: @@ -978,14 +978,14 @@ subsection \Lebesgue measurable sets\ -abbreviation%important lmeasurable :: "'a::euclidean_space set set" +abbreviation\<^marker>\tag important\ lmeasurable :: "'a::euclidean_space set set" where "lmeasurable \ fmeasurable lebesgue" lemma not_measurable_UNIV [simp]: "UNIV \ lmeasurable" by (simp add: fmeasurable_def) -lemma%important lmeasurable_iff_integrable: +lemma\<^marker>\tag important\ lmeasurable_iff_integrable: "S \ lmeasurable \ integrable lebesgue (indicator S :: 'a::euclidean_space \ 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\Translation preserves Lebesgue measure\ +subsection\<^marker>\tag unimportant\\Translation preserves Lebesgue measure\ lemma sigma_sets_image: assumes S: "S \ sigma_sets \ M" and "M \ Pow \" "f ` \ = \" "inj_on f \" @@ -1128,12 +1128,12 @@ lemma summable_iff_suminf_neq_top: "(\n. f n \ 0) \ \ summable f \ (\i. ennreal (f i)) = top" by (metis summable_suminf_not_top) -proposition%important starlike_negligible_bounded_gmeasurable: +proposition\<^marker>\tag important\ starlike_negligible_bounded_gmeasurable: fixes S :: "'a :: euclidean_space set" assumes S: "S \ sets lebesgue" and "bounded S" and eq1: "\c x. \(c *\<^sub>R x) \ S; 0 \ c; x \ S\ \ c = 1" shows "S \ null_sets lebesgue" -proof%unimportant - +proof - obtain M where "0 < M" "S \ ball 0 M" using \bounded S\ by (auto dest: bounded_subset_ballD) @@ -1252,10 +1252,10 @@ qed qed -lemma%important outer_regular_lborel: +lemma\<^marker>\tag important\ outer_regular_lborel: assumes B: "B \ sets borel" and "0 < (e::real)" obtains U where "open U" "B \ U" "emeasure lborel (U - B) < e" -proof%unimportant - +proof - obtain U where U: "open U" "B \ U" and "emeasure lborel (U-B) \ e/2" using outer_regular_lborel_le [OF B, of "e/2"] \e > 0\ by force diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Linear_Algebra.thy --- 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 \More interesting properties of the norm\ +subsection\<^marker>\tag unimportant\ \More interesting properties of the norm\ unbundle inner_syntax @@ -230,7 +230,7 @@ subsection \Orthogonality\ -definition%important (in real_inner) "orthogonal x y \ x \ y = 0" +definition\<^marker>\tag important\ (in real_inner) "orthogonal x y \ x \ y = 0" context real_inner begin @@ -299,9 +299,9 @@ subsection \Orthogonality of a transformation\ -definition%important "orthogonal_transformation f \ linear f \ (\v w. f v \ f w = v \ w)" +definition\<^marker>\tag important\ "orthogonal_transformation f \ linear f \ (\v w. f v \ f w = v \ w)" -lemma%unimportant orthogonal_transformation: +lemma\<^marker>\tag unimportant\ orthogonal_transformation: "orthogonal_transformation f \ linear f \ (\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 (\x. x)" +lemma\<^marker>\tag unimportant\ orthogonal_transformation_id [simp]: "orthogonal_transformation (\x. x)" by (simp add: linear_iff orthogonal_transformation_def) -lemma%unimportant orthogonal_orthogonal_transformation: +lemma\<^marker>\tag unimportant\ orthogonal_orthogonal_transformation: "orthogonal_transformation f \ orthogonal (f x) (f y) \ orthogonal x y" by (simp add: orthogonal_def orthogonal_transformation_def) -lemma%unimportant orthogonal_transformation_compose: +lemma\<^marker>\tag unimportant\ orthogonal_transformation_compose: "\orthogonal_transformation f; orthogonal_transformation g\ \ orthogonal_transformation(f \ g)" by (auto simp: orthogonal_transformation_def linear_compose) -lemma%unimportant orthogonal_transformation_neg: +lemma\<^marker>\tag unimportant\ orthogonal_transformation_neg: "orthogonal_transformation(\x. -(f x)) \ orthogonal_transformation f" by (auto simp: orthogonal_transformation_def dest: linear_compose_neg) -lemma%unimportant orthogonal_transformation_scaleR: "orthogonal_transformation f \ f (c *\<^sub>R v) = c *\<^sub>R f v" +lemma\<^marker>\tag unimportant\ orthogonal_transformation_scaleR: "orthogonal_transformation f \ 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>\tag unimportant\ orthogonal_transformation_linear: "orthogonal_transformation f \ linear f" by (simp add: orthogonal_transformation_def) -lemma%unimportant orthogonal_transformation_inj: +lemma\<^marker>\tag unimportant\ orthogonal_transformation_inj: "orthogonal_transformation f \ inj f" unfolding orthogonal_transformation_def inj_on_def by (metis vector_eq) -lemma%unimportant orthogonal_transformation_surj: +lemma\<^marker>\tag unimportant\ orthogonal_transformation_surj: "orthogonal_transformation f \ surj f" for f :: "'a::euclidean_space \ 'a::euclidean_space" by (simp add: linear_injective_imp_surjective orthogonal_transformation_inj orthogonal_transformation_linear) -lemma%unimportant orthogonal_transformation_bij: +lemma\<^marker>\tag unimportant\ orthogonal_transformation_bij: "orthogonal_transformation f \ bij f" for f :: "'a::euclidean_space \ 'a::euclidean_space" by (simp add: bij_def orthogonal_transformation_inj orthogonal_transformation_surj) -lemma%unimportant orthogonal_transformation_inv: +lemma\<^marker>\tag unimportant\ orthogonal_transformation_inv: "orthogonal_transformation f \ orthogonal_transformation (inv f)" for f :: "'a::euclidean_space \ '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>\tag unimportant\ orthogonal_transformation_norm: "orthogonal_transformation f \ norm (f x) = norm x" by (metis orthogonal_transformation) subsection \Bilinear functions\ -definition%important +definition\<^marker>\tag important\ bilinear :: "('a::real_vector \ 'b::real_vector \ 'c::real_vector) \ bool" where "bilinear f \ (\x. linear (\y. f x y)) \ (\y. linear (\x. f x y))" @@ -417,7 +417,7 @@ subsection \Adjoints\ -definition%important adjoint :: "(('a::real_inner) \ ('b::real_inner)) \ 'b \ 'a" where +definition\<^marker>\tag important\ adjoint :: "(('a::real_inner) \ ('b::real_inner)) \ 'b \ 'a" where "adjoint f = (SOME f'. \x y. f x \ y = x \ f' y)" lemma adjoint_unique: @@ -579,7 +579,7 @@ done -subsection%unimportant \Euclidean Spaces as Typeclass\ +subsection\<^marker>\tag unimportant\ \Euclidean Spaces as Typeclass\ lemma independent_Basis: "independent Basis" by (rule independent_Basis) @@ -591,7 +591,7 @@ unfolding span_Basis .. -subsection%unimportant \Linearity and Bilinearity continued\ +subsection\<^marker>\tag unimportant\ \Linearity and Bilinearity continued\ lemma linear_bounded: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" @@ -773,7 +773,7 @@ by (metis linear_imp_has_derivative differentiable_def) -subsection%unimportant \We continue\ +subsection\<^marker>\tag unimportant\ \We continue\ lemma independent_bound: fixes S :: "'a::euclidean_space set" @@ -1037,7 +1037,7 @@ subsection \Infinity norm\ -definition%important "infnorm (x::'a::euclidean_space) = Sup {\x \ b\ |b. b \ Basis}" +definition\<^marker>\tag important\ "infnorm (x::'a::euclidean_space) = Sup {\x \ b\ |b. b \ Basis}" lemma infnorm_set_image: fixes x :: "'a::euclidean_space" @@ -1227,7 +1227,7 @@ subsection \Collinearity\ -definition%important collinear :: "'a::real_vector set \ bool" +definition\<^marker>\tag important\ collinear :: "'a::real_vector set \ bool" where "collinear S \ (\u. \x \ S. \ y \ S. \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>\tag unimportant\ orthogonal_extension_strong: fixes S :: "'a::euclidean_space set" assumes S: "pairwise orthogonal S" obtains U where "U \ (insert 0 S) = {}" "pairwise orthogonal (S \ U)" @@ -1648,7 +1648,7 @@ qed -proposition%unimportant orthogonal_to_subspace_exists_gen: +proposition\<^marker>\tag unimportant\ orthogonal_to_subspace_exists_gen: fixes S :: "'a :: euclidean_space set" assumes "span S \ span T" obtains x where "x \ 0" "x \ span T" "\y. y \ span S \ orthogonal x y" @@ -1694,7 +1694,7 @@ qed qed -corollary%unimportant orthogonal_to_subspace_exists: +corollary\<^marker>\tag unimportant\ orthogonal_to_subspace_exists: fixes S :: "'a :: euclidean_space set" assumes "dim S < DIM('a)" obtains x where "x \ 0" "\y. y \ span S \ orthogonal x y" @@ -1706,7 +1706,7 @@ by (auto simp: span_UNIV) qed -corollary%unimportant orthogonal_to_vector_exists: +corollary\<^marker>\tag unimportant\ orthogonal_to_vector_exists: fixes x :: "'a :: euclidean_space" assumes "2 \ DIM('a)" obtains y where "y \ 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>\tag unimportant\ orthogonal_subspace_decomp_exists: fixes S :: "'a :: euclidean_space set" obtains y z where "y \ span S" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Lipschitz.thy --- 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>\tag important\ lipschitz_on where "lipschitz_on C U f \ (0 \ C \ (\x \ U. \y\U. dist (f x) (f y) \ C * dist x y))" bundle lipschitz_syntax begin -notation%important lipschitz_on ("_-lipschitz'_on" [1000]) +notation\<^marker>\tag important\ 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 \Structural introduction rules\ +subsubsection\<^marker>\tag unimportant\ \Structural introduction rules\ named_theorems lipschitz_intros "structural introduction rules for Lipschitz controls" @@ -481,7 +481,7 @@ subsection \Local Lipschitz continuity (uniform for a family of functions)\ -definition%important local_lipschitz:: +definition\<^marker>\tag important\ local_lipschitz:: "'a::metric_space set \ 'b::metric_space set \ ('a \ 'b \ 'c::metric_space) \ bool" where "local_lipschitz T X f \ \x \ X. \t \ T. diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Measurable.thy --- 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 = \ Scan.lift (Scan.succeed (METHOD o Measurable.measurable_tac)) \ +method_setup\<^marker>\tag important\ measurable = \ Scan.lift (Scan.succeed (METHOD o Measurable.measurable_tac)) \ "measurability prover" -simproc_setup%important measurable ("A \ sets M" | "f \ measurable M N") = \K Measurable.simproc\ +simproc_setup\<^marker>\tag important\ measurable ("A \ sets M" | "f \ measurable M N") = \K Measurable.simproc\ setup \ Global_Theory.add_thms_dynamic (\<^binding>\measurable\, 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 \Measurability for (co)inductive predicates\ +subsection\<^marker>\tag unimportant\ \Measurability for (co)inductive predicates\ lemma measurable_bot[measurable]: "bot \ measurable M (count_space UNIV)" by (simp add: bot_fun_def) diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Measure_Space.thy --- 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>\tag unimportant\ "Relate extended reals and the indicator function" lemma suminf_cmult_indicator: fixes f :: "nat \ ennreal" @@ -59,7 +59,7 @@ is also used to represent sigma algebras (with an arbitrary emeasure). \ -subsection%unimportant "Extend binary sets" +subsection\<^marker>\tag unimportant\ "Extend binary sets" lemma LIMSEQ_binaryset: assumes f: "f {} = 0" @@ -91,7 +91,7 @@ shows "f {} = 0 \ (\n. f (binaryset A B n)) = f A + f B" by (metis binaryset_sums sums_unique) -subsection%unimportant \Properties of a premeasure \<^term>\\\\ +subsection\<^marker>\tag unimportant\ \Properties of a premeasure \<^term>\\\\ text \ The definitions for \<^const>\positive\ and \<^const>\countably_additive\ 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 \Properties of \<^const>\emeasure\\ +subsection\<^marker>\tag unimportant\ \Properties of \<^const>\emeasure\\ 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 \\\\-null sets\ -definition%important null_sets :: "'a measure \ 'a set set" where +definition\<^marker>\tag important\ null_sets :: "'a measure \ 'a set set" where "null_sets M = {N\sets M. emeasure M N = 0}" lemma null_setsD1[dest]: "A \ null_sets M \ emeasure M A = 0" @@ -989,7 +989,7 @@ subsection \The almost everywhere filter (i.e.\ quantifier)\ -definition%important ae_filter :: "'a measure \ 'a filter" where +definition\<^marker>\tag important\ ae_filter :: "'a measure \ 'a filter" where "ae_filter M = (INF N\null_sets M. principal (space M - N))" abbreviation almost_everywhere :: "'a measure \ ('a \ bool) \ bool" where @@ -1265,7 +1265,7 @@ subsection \\\\-finite Measures\ -locale%important sigma_finite_measure = +locale\<^marker>\tag important\ sigma_finite_measure = fixes M :: "'a measure" assumes sigma_finite_countable: "\A::'a set set. countable A \ A \ sets M \ (\A) = space M \ (\a\A. emeasure M a \ \)" @@ -1387,7 +1387,7 @@ subsection \Measure space induced by distribution of \<^const>\measurable\-functions\ -definition%important distr :: "'a measure \ 'b measure \ ('a \ 'b) \ 'b measure" where +definition\<^marker>\tag important\ distr :: "'a measure \ 'b measure \ ('a \ 'b) \ 'b measure" where "distr M N f = measure_of (space N) (sets N) (\A. emeasure M (f -` A \ 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 \Real measure values\ +subsection\<^marker>\tag unimportant\ \Real measure values\ lemma ring_of_finite_sets: "ring_of_sets (space M) {A\sets M. emeasure M A \ top}" proof (rule ring_of_setsI) @@ -1709,7 +1709,7 @@ subsection \Set of measurable sets with finite measure\ -definition%important fmeasurable :: "'a measure \ 'a set set" where +definition\<^marker>\tag important\ fmeasurable :: "'a measure \ 'a set set" where "fmeasurable M = {A\sets M. emeasure M A < \}" lemma fmeasurableD[dest, measurable_dest]: "A \ fmeasurable M \ A \ sets M" @@ -1749,7 +1749,7 @@ using emeasure_mono[of "a - b" a M] * by (auto simp: fmeasurable_def) qed -subsection%unimportant\Measurable sets formed by unions and intersections\ +subsection\<^marker>\tag unimportant\\Measurable sets formed by unions and intersections\ lemma fmeasurable_Diff: "A \ fmeasurable M \ B \ sets M \ A - B \ fmeasurable M" using fmeasurableI2[of A M "A - B"] by auto @@ -2073,7 +2073,7 @@ subsection \Measure spaces with \<^term>\emeasure M (space M) < \\\ -locale%important finite_measure = sigma_finite_measure M for M + +locale\<^marker>\tag important\ finite_measure = sigma_finite_measure M for M + assumes finite_emeasure_space: "emeasure M (space M) \ 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 \Counting space\ +subsection\<^marker>\tag unimportant\ \Counting space\ lemma strict_monoI_Suc: assumes ord [simp]: "(\n. f n < f (Suc n))" shows "strict_mono f" @@ -2444,7 +2444,7 @@ show "sigma_finite_measure (count_space A)" .. qed -subsection%unimportant \Measure restricted to space\ +subsection\<^marker>\tag unimportant\ \Measure restricted to space\ lemma emeasure_restrict_space: assumes "\ \ space M \ sets M" "A \ \" @@ -2603,7 +2603,7 @@ finally show "emeasure M X = emeasure N X" . qed fact -subsection%unimportant \Null measure\ +subsection\<^marker>\tag unimportant\ \Null measure\ definition null_measure :: "'a measure \ 'a measure" where "null_measure M = sigma (space M) (sets M)" @@ -2627,7 +2627,7 @@ subsection \Scaling a measure\ -definition%important scale_measure :: "ennreal \ 'a measure \ 'a measure" where +definition\<^marker>\tag important\ scale_measure :: "ennreal \ 'a measure \ 'a measure" where "scale_measure r M = measure_of (space M) (sets M) (\A. r * emeasure M A)" lemma space_scale_measure: "space (scale_measure r M) = space M" @@ -2829,7 +2829,7 @@ done qed -text%important \ +text\<^marker>\tag important\ \ Define a lexicographical order on \<^type>\measure\, in the order space, sets and measure. The parts of the lexicographical order are point-wise ordered. \ @@ -2847,10 +2847,10 @@ if sets M = sets N then emeasure M \ emeasure N else sets M \ sets N else space M \ space N)" by (auto elim: less_eq_measure.cases intro: less_eq_measure.intros) -definition%important less_measure :: "'a measure \ 'a measure \ bool" where +definition\<^marker>\tag important\ less_measure :: "'a measure \ 'a measure \ bool" where "less_measure M N \ (M \ N \ \ N \ M)" -definition%important bot_measure :: "'a measure" where +definition\<^marker>\tag important\ bot_measure :: "'a measure" where "bot_measure = sigma {} {}" lemma @@ -2874,7 +2874,7 @@ by (cases "X \ sets M") (auto simp: emeasure_notin_sets) done -definition%important sup_measure' :: "'a measure \ 'a measure \ 'a measure" where +definition\<^marker>\tag important\ sup_measure' :: "'a measure \ 'a measure \ 'a measure" where "sup_measure' A B = measure_of (space A) (sets A) (\X. SUP Y\sets A. emeasure A (X \ Y) + emeasure B (X \ - Y))" @@ -3008,7 +3008,7 @@ qed qed simp_all -definition%important sup_lexord :: "'a \ 'a \ ('a \ 'b::order) \ 'a \ 'a \ 'a" where +definition\<^marker>\tag important\ sup_lexord :: "'a \ 'a \ ('a \ 'b::order) \ 'a \ 'a \ 'a" where "sup_lexord A B k s c = (if k A = k B then c else if \ k A \ k B \ \ k B \ k A then s else @@ -3038,7 +3038,7 @@ instantiation measure :: (type) semilattice_sup begin -definition%important sup_measure :: "'a measure \ 'a measure \ 'a measure" where +definition\<^marker>\tag important\ sup_measure :: "'a measure \ 'a measure \ 'a measure" where "sup_measure A B = sup_lexord A B space (sigma (space A \ space B) {}) (sup_lexord A B sets (sigma (space A) (sets A \ sets B)) (sup_measure' A B))" @@ -3141,7 +3141,7 @@ lemma UN_space_closed: "\(sets ` S) \ Pow (\(space ` S))" using sets.space_closed by auto -definition%important +definition\<^marker>\tag important\ Sup_lexord :: "('a \ 'b::complete_lattice) \ ('a set \ 'a) \ ('a set \ 'a) \ 'a set \ 'a" where "Sup_lexord k c s A = @@ -3203,7 +3203,7 @@ "finite I \ I \ {} \ (\i. i \ I \ sets i = sets M) \ 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 \ 'a measure" where +definition\<^marker>\tag important\ Sup_measure' :: "'a measure set \ 'a measure" where "Sup_measure' M = measure_of (\a\M. space a) (\a\M. sets a) (\X. (SUP P\{P. finite P \ P \ 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 \ 'a measure" where +definition\<^marker>\tag important\ Sup_measure :: "'a measure set \ 'a measure" where "Sup_measure = Sup_lexord space (Sup_lexord sets Sup_measure' (\U. sigma (\u\U. space u) (\u\U. sets u))) (\U. sigma (\u\U. space u) {})" -definition%important Inf_measure :: "'a measure set \ 'a measure" where +definition\<^marker>\tag important\ Inf_measure :: "'a measure set \ 'a measure" where "Inf_measure A = Sup {x. \a\A. x \ a}" -definition%important inf_measure :: "'a measure \ 'a measure \ 'a measure" where +definition\<^marker>\tag important\ inf_measure :: "'a measure \ 'a measure \ 'a measure" where "inf_measure a b = Inf {a, b}" -definition%important top_measure :: "'a measure" where +definition\<^marker>\tag important\ top_measure :: "'a measure" where "top_measure = Inf {}" instance @@ -3527,7 +3527,7 @@ qed qed -subsubsection%unimportant \Supremum of a set of \\\-algebras\ +subsubsection\<^marker>\tag unimportant\ \Supremum of a set of \\\-algebras\ lemma space_Sup_eq_UN: "space (Sup M) = (\x\M. space x)" unfolding Sup_measure_def diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy --- 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 \Approximating functions\ +subsection\<^marker>\tag unimportant\ \Approximating functions\ lemma AE_upper_bound_inf_ennreal: fixes F G::"'a \ ennreal" @@ -115,7 +115,7 @@ \ -definition%important "simple_function M g \ +definition\<^marker>\tag important\ "simple_function M g \ finite (g ` space M) \ (\x \ g ` space M. g -` {x} \ space M \ sets M)" @@ -301,11 +301,11 @@ shows "simple_function M (\x. real (f x))" by (rule simple_function_compose1[OF sf]) -lemma%important borel_measurable_implies_simple_function_sequence: +lemma\<^marker>\tag important\ borel_measurable_implies_simple_function_sequence: fixes u :: "'a \ ennreal" assumes u[measurable]: "u \ borel_measurable M" shows "\f. incseq f \ (\i. (\x. f i x < top) \ simple_function M (f i)) \ 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>\tag important\ simple_function_induct [consumes 1, case_names cong set mult add, induct set: simple_function]: fixes u :: "'a \ ennreal" assumes u: "simple_function M u" @@ -406,7 +406,7 @@ assumes mult: "\u c. P u \ P (\x. c * u x)" assumes add: "\u v. P u \ P v \ P (\x. v x + u x)" shows "P u" -proof%unimportant (rule cong) +proof (rule cong) from AE_space show "AE x in M. (\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" proof eventually_elim fix x assume x: "x \ space M" @@ -468,7 +468,7 @@ qed fact qed -lemma%important borel_measurable_induct +lemma\<^marker>\tag important\ borel_measurable_induct [consumes 1, case_names cong set mult add seq, induct set: borel_measurable]: fixes u :: "'a \ ennreal" assumes u: "u \ borel_measurable M" @@ -478,8 +478,8 @@ assumes add: "\u v. u \ borel_measurable M\ (\x. x \ space M \ u x < top) \ P u \ v \ borel_measurable M \ (\x. x \ space M \ v x < top) \ (\x. x \ space M \ u x = 0 \ v x = 0) \ P v \ P (\x. v x + u x)" assumes seq: "\U. (\i. U i \ borel_measurable M) \ (\i x. x \ space M \ U i x < top) \ (\i. P (U i)) \ incseq U \ u = (SUP i. U i) \ 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: "\i. simple_function M (U i)" "incseq U" "\i x. U i x < top" and sup: "\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 \ ('a \ ennreal) \ ennreal" ("integral\<^sup>S") where +definition\<^marker>\tag important\ simple_integral :: "'a measure \ ('a \ ennreal) \ ennreal" ("integral\<^sup>S") where "integral\<^sup>S M f = (\x \ f ` space M. x * emeasure M (f -` {x} \ space M))" syntax @@ -813,7 +813,7 @@ subsection \Integral on nonnegative functions\ -definition%important nn_integral :: "'a measure \ ('a \ ennreal) \ ennreal" ("integral\<^sup>N") where +definition\<^marker>\tag important\ nn_integral :: "'a measure \ ('a \ ennreal) \ ennreal" ("integral\<^sup>N") where "integral\<^sup>N M f = (SUP g \ {g. simple_function M g \ g \ 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 \Distributions\ +subsubsection\<^marker>\tag unimportant\ \Distributions\ lemma nn_integral_distr: assumes T: "T \ measurable M M'" and f: "f \ 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 \Counting space\ +subsubsection\<^marker>\tag unimportant\ \Counting space\ lemma simple_function_count_space[simp]: "simple_function (count_space A) f \ finite (f ` A)" @@ -2083,7 +2083,7 @@ subsubsection \Measure spaces with an associated density\ -definition%important density :: "'a measure \ ('a \ ennreal) \ 'a measure" where +definition\<^marker>\tag important\ density :: "'a measure \ ('a \ ennreal) \ 'a measure" where "density M f = measure_of (space M) (sets M) (\A. \\<^sup>+ x. f x * indicator A x \M)" lemma @@ -2170,11 +2170,11 @@ by (intro exI[of _ "N \ {x\space M. f x = 0}"] conjI *) (auto elim: eventually_elim2) qed -lemma%important nn_integral_density: +lemma\<^marker>\tag important\ nn_integral_density: assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" shows "integral\<^sup>N (density M f) g = (\\<^sup>+ x. f x * g x \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 \Point measure\ -definition%important point_measure :: "'a set \ ('a \ ennreal) \ 'a measure" where +definition\<^marker>\tag important\ point_measure :: "'a set \ ('a \ ennreal) \ 'a measure" where "point_measure A f = density (count_space A) f" lemma @@ -2359,7 +2359,7 @@ subsubsection \Uniform measure\ -definition%important "uniform_measure M A = density M (\x. indicator A x / emeasure M A)" +definition\<^marker>\tag important\ "uniform_measure M A = density M (\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 \Null measure\ +subsubsection\<^marker>\tag unimportant\ \Null measure\ lemma null_measure_eq_density: "null_measure M = density M (\_. 0)" by (intro measure_eqI) (simp_all add: emeasure_density) @@ -2451,7 +2451,7 @@ subsubsection \Uniform count measure\ -definition%important "uniform_count_measure A = point_measure A (\x. 1 / card A)" +definition\<^marker>\tag important\ "uniform_count_measure A = point_measure A (\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) \ True" by(simp_all add: sets_uniform_count_measure) -subsubsection%unimportant \Scaled measure\ +subsubsection\<^marker>\tag unimportant\ \Scaled measure\ lemma nn_integral_scale_measure: assumes f: "f \ borel_measurable M" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Norm_Arith.thy --- 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 \normarith.ML\ -method_setup%important norm = \ +method_setup\<^marker>\tag important\ norm = \ Scan.succeed (SIMPLE_METHOD' o NormArith.norm_arith_tac) \ "prove simple linear statements about vector norms" text \Hence more metric properties.\ -text%important \%whitespace\ +text\<^marker>\tag important\ \%whitespace\ proposition dist_triangle_add: fixes x y x' y' :: "'a::real_normed_vector" shows "dist (x + y) (x' + y') \ dist x x' + dist y y'" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Operator_Norm.thy --- 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 \This formulation yields zero if \'a\ is the trivial vector space.\ -text%important \%whitespace\ -definition%important +text\<^marker>\tag important\ \%whitespace\ +definition\<^marker>\tag important\ onorm :: "('a::real_normed_vector \ 'b::real_normed_vector) \ real" where "onorm f = (SUP x. norm (f x) / norm x)" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Ordered_Euclidean_Space.thy --- 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 \ Basis \ fst i = 0 \ snd i \ Basis \ snd i = 0 \ fst i \ Basis" by (cases i) (auto simp: Basis_prod_def) -instantiation%unimportant prod :: (abs, abs) abs +instantiation\<^marker>\tag unimportant\ prod :: (abs, abs) abs begin definition "\x\ = (\fst x\, \snd x\)" @@ -299,11 +299,11 @@ instantiation vec :: (ordered_euclidean_space, finite) ordered_euclidean_space begin -definition%important "inf x y = (\ i. inf (x $ i) (y $ i))" -definition%important "sup x y = (\ i. sup (x $ i) (y $ i))" -definition%important "Inf X = (\ i. (INF x\X. x $ i))" -definition%important "Sup X = (\ i. (SUP x\X. x $ i))" -definition%important "\x\ = (\ i. \x $ i\)" +definition\<^marker>\tag important\ "inf x y = (\ i. inf (x $ i) (y $ i))" +definition\<^marker>\tag important\ "sup x y = (\ i. sup (x $ i) (y $ i))" +definition\<^marker>\tag important\ "Inf X = (\ i. (INF x\X. x $ i))" +definition\<^marker>\tag important\ "Sup X = (\ i. (SUP x\X. x $ i))" +definition\<^marker>\tag important\ "\x\ = (\ i. \x $ i\)" instance apply standard diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Path_Connected.thy --- 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 \Paths and Arcs\ -definition%important path :: "(real \ 'a::topological_space) \ bool" +definition\<^marker>\tag important\ path :: "(real \ 'a::topological_space) \ bool" where "path g \ continuous_on {0..1} g" -definition%important pathstart :: "(real \ 'a::topological_space) \ 'a" +definition\<^marker>\tag important\ pathstart :: "(real \ 'a::topological_space) \ 'a" where "pathstart g = g 0" -definition%important pathfinish :: "(real \ 'a::topological_space) \ 'a" +definition\<^marker>\tag important\ pathfinish :: "(real \ 'a::topological_space) \ 'a" where "pathfinish g = g 1" -definition%important path_image :: "(real \ 'a::topological_space) \ 'a set" +definition\<^marker>\tag important\ path_image :: "(real \ 'a::topological_space) \ 'a set" where "path_image g = g ` {0 .. 1}" -definition%important reversepath :: "(real \ 'a::topological_space) \ real \ 'a" +definition\<^marker>\tag important\ reversepath :: "(real \ 'a::topological_space) \ real \ 'a" where "reversepath g = (\x. g(1 - x))" -definition%important joinpaths :: "(real \ 'a::topological_space) \ (real \ 'a) \ real \ 'a" +definition\<^marker>\tag important\ joinpaths :: "(real \ 'a::topological_space) \ (real \ 'a) \ real \ 'a" (infixr "+++" 75) where "g1 +++ g2 = (\x. if x \ 1/2 then g1 (2 * x) else g2 (2 * x - 1))" -definition%important simple_path :: "(real \ 'a::topological_space) \ bool" +definition\<^marker>\tag important\ simple_path :: "(real \ 'a::topological_space) \ bool" where "simple_path g \ path g \ (\x\{0..1}. \y\{0..1}. g x = g y \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0)" -definition%important arc :: "(real \ 'a :: topological_space) \ bool" +definition\<^marker>\tag important\ arc :: "(real \ 'a :: topological_space) \ bool" where "arc g \ path g \ inj_on g {0..1}" -subsection%unimportant\Invariance theorems\ +subsection\<^marker>\tag unimportant\\Invariance theorems\ lemma path_eq: "path p \ (\t. t \ {0..1} \ p t = q t) \ 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\Basic lemmas about paths\ +subsection\<^marker>\tag unimportant\\Basic lemmas about paths\ lemma pathin_iff_path_real [simp]: "pathin euclideanreal g \ path g" by (simp add: pathin_def path_def) @@ -298,7 +298,7 @@ qed -subsection%unimportant \Path Images\ +subsection\<^marker>\tag unimportant\ \Path Images\ lemma bounded_path_image: "path g \ 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\Simple paths with the endpoints removed\ +subsection\<^marker>\tag unimportant\\Simple paths with the endpoints removed\ lemma simple_path_endless: "simple_path c \ path_image c - {pathstart c,pathfinish c} = c ` {0<..<1}" @@ -414,7 +414,7 @@ by (simp add: simple_path_endless) -subsection%unimportant\The operations on paths\ +subsection\<^marker>\tag unimportant\\The operations on paths\ lemma path_image_subset_reversepath: "path_image(reversepath g) \ path_image g" by (auto simp: path_image_def reversepath_def) @@ -562,7 +562,7 @@ by (rule ext) (auto simp: mult.commute) -subsection%unimportant\Some reversed and "if and only if" versions of joining theorems\ +subsection\<^marker>\tag unimportant\\Some reversed and "if and only if" versions of joining theorems\ lemma path_join_path_ends: fixes g1 :: "real \ 'a::metric_space" @@ -695,7 +695,7 @@ using pathfinish_in_path_image by (fastforce simp: arc_join_eq) -subsection%unimportant\The joining of paths is associative\ +subsection\<^marker>\tag unimportant\\The joining of paths is associative\ lemma path_assoc: "\pathfinish p = pathstart q; pathfinish q = pathstart r\ @@ -745,7 +745,7 @@ \ arc(p +++ (q +++ r)) \ arc((p +++ q) +++ r)" by (simp add: arc_simple_path simple_path_assoc) -subsubsection%unimportant\Symmetry and loops\ +subsubsection\<^marker>\tag unimportant\\Symmetry and loops\ lemma path_sym: "\pathfinish p = pathstart q; pathfinish q = pathstart p\ \ path(p +++ q) \ path(q +++ p)" @@ -764,7 +764,7 @@ subsection\Subpath\ -definition%important subpath :: "real \ real \ (real \ 'a) \ real \ 'a::real_normed_vector" +definition\<^marker>\tag important\ subpath :: "real \ real \ (real \ 'a) \ real \ 'a::real_normed_vector" where "subpath a b g \ \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\There is a subpath to the frontier\ +subsection\<^marker>\tag unimportant\\There is a subpath to the frontier\ lemma subpath_to_frontier_explicit: fixes S :: "'a::metric_space set" @@ -1059,7 +1059,7 @@ subsection \Shift Path to Start at Some Given Point\ -definition%important shiftpath :: "real \ (real \ 'a::topological_space) \ real \ 'a" +definition\<^marker>\tag important\ shiftpath :: "real \ (real \ 'a::topological_space) \ real \ 'a" where "shiftpath a f = (\x. if (a + x) \ 1 then f (a + x) else f (a + x - 1))" lemma pathstart_shiftpath: "a \ 1 \ pathstart (shiftpath a g) = g a" @@ -1171,7 +1171,7 @@ subsection \Straight-Line Paths\ -definition%important linepath :: "'a::real_normed_vector \ 'a \ real \ 'a" +definition\<^marker>\tag important\ linepath :: "'a::real_normed_vector \ 'a \ real \ 'a" where "linepath a b = (\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\Segments via convex hulls\ +subsection\<^marker>\tag unimportant\\Segments via convex hulls\ lemma segments_subset_convex_hull: "closed_segment a b \ (convex hull {a,b,c})" @@ -1391,7 +1391,7 @@ qed -subsection%unimportant \Bounding a point away from a path\ +subsection\<^marker>\tag unimportant\ \Bounding a point away from a path\ lemma not_on_path_ball: fixes g :: "real \ 'a::heine_borel" @@ -1425,10 +1425,10 @@ text \Original formalization by Tom Hales\ -definition%important "path_component s x y \ +definition\<^marker>\tag important\ "path_component s x y \ (\g. path g \ path_image g \ s \ pathstart g = x \ pathfinish g = y)" -abbreviation%important +abbreviation\<^marker>\tag important\ "path_component_set s x \ 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 \Path components as sets\ +subsubsection\<^marker>\tag unimportant\ \Path components as sets\ lemma path_component_set: "path_component_set s x = @@ -1502,7 +1502,7 @@ subsection \Path connectedness of a space\ -definition%important "path_connected s \ +definition\<^marker>\tag important\ "path_connected s \ (\x\s. \y\s. \g. path g \ path_image g \ s \ pathstart g = x \ pathfinish g = y)" lemma path_connectedin_iff_path_connected_real [simp]: @@ -1814,7 +1814,7 @@ qed -subsection%unimportant\Lemmas about path-connectedness\ +subsection\<^marker>\tag unimportant\\Lemmas about path-connectedness\ lemma path_connected_linear_image: fixes f :: "'a::real_normed_vector \ '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\Path components\ +subsection\<^marker>\tag unimportant\\Path components\ lemma Union_path_component [simp]: "Union {path_component_set S x |x. x \ S} = S" @@ -2736,7 +2736,7 @@ by (auto simp: path_connected_annulus [OF assms] path_connected_imp_connected) -subsection%unimportant\Relations between components and path components\ +subsection\<^marker>\tag unimportant\\Relations between components and path components\ lemma open_connected_component: fixes s :: "'a::real_normed_vector set" @@ -2841,7 +2841,7 @@ qed -subsection%unimportant\Existence of unbounded components\ +subsection\<^marker>\tag unimportant\\Existence of unbounded components\ lemma cobounded_unbounded_component: fixes s :: "'a :: euclidean_space set" @@ -2929,13 +2929,13 @@ subsection\The \inside\ and \outside\ of a Set\ -text%important\The inside comprises the points in a bounded connected component of the set's complement. +text\<^marker>\tag important\\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.\ -definition%important inside where +definition\<^marker>\tag important\ inside where "inside S \ {x. (x \ S) \ bounded(connected_component_set ( - S) x)}" -definition%important outside where +definition\<^marker>\tag important\ outside where "outside S \ -S \ {x. \ bounded(connected_component_set (- S) x)}" lemma outside: "outside S = {x. \ bounded(connected_component_set (- S) x)}" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Polytope.thy --- 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 \Faces of a (usually convex) set\ -definition%important face_of :: "['a::real_vector set, 'a set] \ bool" (infixr "(face'_of)" 50) +definition\<^marker>\tag important\ face_of :: "['a::real_vector set, 'a set] \ bool" (infixr "(face'_of)" 50) where "T face_of S \ T \ S \ convex T \ @@ -645,7 +645,7 @@ text\That is, faces that are intersection with supporting hyperplane\ -definition%important exposed_face_of :: "['a::euclidean_space set, 'a set] \ bool" +definition\<^marker>\tag important\ exposed_face_of :: "['a::euclidean_space set, 'a set] \ bool" (infixr "(exposed'_face'_of)" 50) where "T exposed_face_of S \ T face_of S \ (\a b. S \ {x. a \ x \ b} \ T = S \ {x. a \ x = b})" @@ -876,7 +876,7 @@ subsection\Extreme points of a set: its singleton faces\ -definition%important extreme_point_of :: "['a::real_vector, 'a set] \ bool" +definition\<^marker>\tag important\ extreme_point_of :: "['a::real_vector, 'a set] \ bool" (infixr "(extreme'_point'_of)" 50) where "x extreme_point_of S \ x \ S \ (\a \ S. \b \ S. x \ open_segment a b)" @@ -976,7 +976,7 @@ subsection\Facets\ -definition%important facet_of :: "['a::euclidean_space set, 'a set] \ bool" +definition\<^marker>\tag important\ facet_of :: "['a::euclidean_space set, 'a set] \ bool" (infixr "(facet'_of)" 50) where "F facet_of S \ F face_of S \ F \ {} \ aff_dim F = aff_dim S - 1" @@ -1022,7 +1022,7 @@ subsection \Edges: faces of affine dimension 1\ (*FIXME too small subsection, rearrange? *) -definition%important edge_of :: "['a::euclidean_space set, 'a set] \ bool" (infixr "(edge'_of)" 50) +definition\<^marker>\tag important\ edge_of :: "['a::euclidean_space set, 'a set] \ bool" (infixr "(edge'_of)" 50) where "e edge_of S \ e face_of S \ aff_dim e = 1" lemma edge_of_imp_subset: @@ -1601,7 +1601,7 @@ subsection\Polytopes\ -definition%important polytope where +definition\<^marker>\tag important\ polytope where "polytope S \ \v. finite v \ S = convex hull v" lemma polytope_translation_eq: "polytope (image (\x. a + x) S) \ polytope S" @@ -1696,7 +1696,7 @@ subsection\Polyhedra\ -definition%important polyhedron where +definition\<^marker>\tag important\ polyhedron where "polyhedron S \ \F. finite F \ S = \ F \ @@ -3088,7 +3088,7 @@ text\The notion of n-simplex for integer \<^term>\n \ -1\\ -definition%important simplex :: "int \ 'a::euclidean_space set \ bool" (infix "simplex" 50) +definition\<^marker>\tag important\ simplex :: "int \ 'a::euclidean_space set \ bool" (infix "simplex" 50) where "n simplex S \ \C. \ affine_dependent C \ int(card C) = n + 1 \ S = convex hull C" lemma simplex: @@ -3211,7 +3211,7 @@ subsection \Simplicial complexes and triangulations\ -definition%important simplicial_complex where +definition\<^marker>\tag important\ simplicial_complex where "simplicial_complex \ \ finite \ \ (\S \ \. \n. n simplex S) \ @@ -3219,7 +3219,7 @@ (\S S'. S \ \ \ S' \ \ \ (S \ S') face_of S \ (S \ S') face_of S')" -definition%important triangulation where +definition\<^marker>\tag important\ triangulation where "triangulation \ \ finite \ \ (\T \ \. \n. n simplex T) \ diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Product_Vector.thy --- 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 \ 'b \ 'c \ 'b \ '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>\tag important\ scale_prod: "scale x (a, b) = (s1 x a, s2 x b)" by (auto simp: scale_def) -sublocale%important p: module scale +sublocale\<^marker>\tag important\ 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>\tag unimportant\ 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>\tag unimportant\ prod :: (metric_space, metric_space) uniformity_dist begin definition [code del]: @@ -253,8 +253,8 @@ subsection \Product is a Complete Metric Space\ -instance%important prod :: (complete_space, complete_space) complete_space -proof%unimportant +instance\<^marker>\tag important\ prod :: (complete_space, complete_space) complete_space +proof fix X :: "nat \ 'a \ 'b" assume "Cauchy X" have 1: "(\n. fst (X n)) \ lim (\n. fst (X n))" using Cauchy_fst [OF \Cauchy X\] @@ -310,9 +310,9 @@ declare [[code abort: "norm::('a::real_normed_vector*'b::real_normed_vector) \ real"]] -instance%important prod :: (banach, banach) banach .. +instance\<^marker>\tag important\ prod :: (banach, banach) banach .. -subsubsection%unimportant \Pair operations are linear\ +subsubsection\<^marker>\tag unimportant\ \Pair operations are linear\ lemma bounded_linear_fst: "bounded_linear fst" using fst_add fst_scaleR @@ -352,7 +352,7 @@ then show "\K. \x. norm (f x, g x) \ norm x * K" .. qed -subsubsection%unimportant \Frechet derivatives involving pairs\ +subsubsection\<^marker>\tag unimportant\ \Frechet derivatives involving pairs\ 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 \Vector derivatives involving pairs\ +subsubsection\<^marker>\tag unimportant\ \Vector derivatives involving pairs\ lemma has_vector_derivative_Pair[derivative_intros]: assumes "(f has_vector_derivative f') (at x within s)" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Radon_Nikodym.thy --- 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 \ 'a measure \ 'a measure" +definition\<^marker>\tag important\ diff_measure :: "'a measure \ 'a measure \ 'a measure" where "diff_measure M N = measure_of (space M) (sets M) (\A. emeasure M A - emeasure N A)" @@ -165,7 +165,7 @@ subsection "Absolutely continuous" -definition%important absolutely_continuous :: "'a measure \ 'a measure \ bool" where +definition\<^marker>\tag important\ absolutely_continuous :: "'a measure \ 'a measure \ bool" where "absolutely_continuous M N \ null_sets M \ null_sets N" lemma absolutely_continuousI_count_space: "absolutely_continuous (count_space A) M" @@ -889,7 +889,7 @@ subsection \Radon-Nikodym derivative\ -definition%important RN_deriv :: "'a measure \ 'a measure \ 'a \ ennreal" where +definition\<^marker>\tag important\ RN_deriv :: "'a measure \ 'a measure \ 'a \ ennreal" where "RN_deriv M N = (if \f. f \ borel_measurable M \ density M f = N then SOME f. f \ borel_measurable M \ density M f = N diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Riemann_Mapping.thy --- 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\Moebius functions are biholomorphisms of the unit disc\ -definition%important Moebius_function :: "[real,complex,complex] \ complex" where +definition\<^marker>\tag important\ Moebius_function :: "[real,complex,complex] \ complex" where "Moebius_function \ \t w z. exp(\ * of_real t) * (z - w) / (1 - cnj w * z)" lemma Moebius_function_simple: @@ -1411,7 +1411,7 @@ qed -subsection%unimportant \More Borsukian results\ +subsection\<^marker>\tag unimportant\ \More Borsukian results\ lemma Borsukian_componentwise_eq: fixes S :: "'a::euclidean_space set" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Set_Integral.thy --- 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 \ (\x. indicator A x *\<^sub>R f x) \ borel_measurable M" +definition\<^marker>\tag important\ "set_borel_measurable M A f \ (\x. indicator A x *\<^sub>R f x) \ borel_measurable M" -definition%important "set_integrable M A f \ integrable M (\x. indicator A x *\<^sub>R f x)" +definition\<^marker>\tag important\ "set_integrable M A f \ integrable M (\x. indicator A x *\<^sub>R f x)" -definition%important "set_lebesgue_integral M A f \ lebesgue_integral M (\x. indicator A x *\<^sub>R f x)" +definition\<^marker>\tag important\ "set_lebesgue_integral M A f \ lebesgue_integral M (\x. indicator A x *\<^sub>R f x)" syntax "_ascii_set_lebesgue_integral" :: "pttrn \ 'a set \ 'a measure \ real \ real" @@ -417,7 +417,7 @@ and intgbl: "set_integrable M (\i. A i) f" shows "LINT x:(\i. A i)|M. f x = (\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 (\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 diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Sigma_Algebra.thy --- 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 \Families of sets\ -locale%important subset_class = +locale\<^marker>\tag important\ subset_class = fixes \ :: "'a set" and M :: "'a set set" assumes space_closed: "M \ Pow \" @@ -39,7 +39,7 @@ subsubsection \Semiring of sets\ -locale%important semiring_of_sets = subset_class + +locale\<^marker>\tag important\ semiring_of_sets = subset_class + assumes empty_sets[iff]: "{} \ M" assumes Int[intro]: "\a b. a \ M \ b \ M \ a \ b \ M" assumes Diff_cover: @@ -76,7 +76,7 @@ subsubsection \Ring of sets\ -locale%important ring_of_sets = semiring_of_sets + +locale\<^marker>\tag important\ ring_of_sets = semiring_of_sets + assumes Un [intro]: "\a b. a \ M \ b \ M \ a \ b \ M" lemma (in ring_of_sets) finite_Union [intro]: @@ -142,7 +142,7 @@ subsubsection \Algebra of sets\ -locale%important algebra = ring_of_sets + +locale\<^marker>\tag important\ algebra = ring_of_sets + assumes top [iff]: "\ \ M" lemma (in algebra) compl_sets [intro]: @@ -221,7 +221,7 @@ "X \ S \ algebra S { {}, X, S - X, S }" by (auto simp: algebra_iff_Int) -subsubsection%unimportant \Restricted algebras\ +subsubsection\<^marker>\tag unimportant\ \Restricted algebras\ abbreviation (in algebra) "restricted_space A \ ((\) A) ` M" @@ -232,7 +232,7 @@ subsubsection \Sigma Algebras\ -locale%important sigma_algebra = algebra + +locale\<^marker>\tag important\ sigma_algebra = algebra + assumes countable_nat_UN [intro]: "\A. range A \ M \ (\i::nat. A i) \ 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 \X \ S\]] by simp -subsubsection%unimportant \Binary Unions\ +subsubsection\<^marker>\tag unimportant\ \Binary Unions\ definition binary :: "'a \ 'a \ nat \ 'a" where "binary a b = (\x. b)(0 := a)" @@ -473,10 +473,10 @@ subsubsection \Initial Sigma Algebra\ -text%important \Sigma algebras can naturally be created as the closure of any set of +text\<^marker>\tag important\ \Sigma algebras can naturally be created as the closure of any set of M with regard to the properties just postulated.\ -inductive_set%important sigma_sets :: "'a set \ 'a set set \ 'a set set" +inductive_set\<^marker>\tag important\ sigma_sets :: "'a set \ 'a set set \ 'a set set" for sp :: "'a set" and A :: "'a set set" where Basic[intro, simp]: "a \ A \ a \ sigma_sets sp A" @@ -819,7 +819,7 @@ thus "(\i::nat. A i) \ M" by (simp add: UN_disjointed_eq) qed -subsubsection%unimportant \Ring generated by a semiring\ +subsubsection\<^marker>\tag unimportant\ \Ring generated by a semiring\ definition (in semiring_of_sets) generated_ring :: "'a set set" where "generated_ring = { \C | C. C \ M \ finite C \ 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 \A Two-Element Series\ +subsubsection\<^marker>\tag unimportant\ \A Two-Element Series\ definition binaryset :: "'a set \ 'a set \ nat \ 'a set" where "binaryset A B = (\x. {})(0 := A, Suc 0 := B)" @@ -966,7 +966,7 @@ subsubsection \Closed CDI\ -definition%important closed_cdi :: "'a set \ 'a set set \ bool" where +definition\<^marker>\tag important\ closed_cdi :: "'a set \ 'a set set \ bool" where "closed_cdi \ M \ M \ Pow \ & (\s \ M. \ - s \ M) & @@ -1200,7 +1200,7 @@ subsubsection \Dynkin systems\ -locale%important Dynkin_system = subset_class + +locale\<^marker>\tag important\ Dynkin_system = subset_class + assumes space: "\ \ M" and compl[intro!]: "\A. A \ M \ \ - A \ M" and UN[intro!]: "\A. disjoint_family A \ range A \ M @@ -1262,7 +1262,7 @@ subsubsection "Intersection sets systems" -definition%important Int_stable :: "'a set set \ bool" where +definition\<^marker>\tag important\ Int_stable :: "'a set set \ bool" where "Int_stable M \ (\ a \ M. \ b \ M. a \ b \ M)" lemma (in algebra) Int_stable: "Int_stable M" @@ -1303,7 +1303,7 @@ subsubsection "Smallest Dynkin systems" -definition%important Dynkin :: "'a set \ 'a set set \ 'a set set" where +definition\<^marker>\tag important\ Dynkin :: "'a set \ 'a set set \ 'a set set" where "Dynkin \ M = (\{D. Dynkin_system \ D \ M \ D})" lemma Dynkin_system_Dynkin: @@ -1450,7 +1450,7 @@ subsubsection \Induction rule for intersection-stable generators\ -text%important \The reason to introduce Dynkin-systems is the following induction rules for \\\-algebras +text\<^marker>\tag important\ \The reason to introduce Dynkin-systems is the following induction rules for \\\-algebras generated by a generator closed under intersection.\ proposition sigma_sets_induct_disjoint[consumes 3, case_names basic empty compl union]: @@ -1476,37 +1476,37 @@ subsection \Measure type\ -definition%important positive :: "'a set set \ ('a set \ ennreal) \ bool" where +definition\<^marker>\tag important\ positive :: "'a set set \ ('a set \ ennreal) \ bool" where "positive M \ \ \ {} = 0" -definition%important countably_additive :: "'a set set \ ('a set \ ennreal) \ bool" where +definition\<^marker>\tag important\ countably_additive :: "'a set set \ ('a set \ ennreal) \ bool" where "countably_additive M f \ (\A. range A \ M \ disjoint_family A \ (\i. A i) \ M \ (\i. f (A i)) = f (\i. A i))" -definition%important measure_space :: "'a set \ 'a set set \ ('a set \ ennreal) \ bool" where +definition\<^marker>\tag important\ measure_space :: "'a set \ 'a set set \ ('a set \ ennreal) \ bool" where "measure_space \ A \ \ sigma_algebra \ A \ positive A \ \ countably_additive A \" -typedef%important 'a measure = +typedef\<^marker>\tag important\ 'a measure = "{(\::'a set, A, \). (\a\-A. \ a = 0) \ measure_space \ A \ }" -proof%unimportant +proof have "sigma_algebra UNIV {{}, UNIV}" by (auto simp: sigma_algebra_iff2) then show "(UNIV, {{}, UNIV}, \A. 0) \ {(\, A, \). (\a\-A. \ a = 0) \ measure_space \ A \} " by (auto simp: measure_space_def positive_def countably_additive_def) qed -definition%important space :: "'a measure \ 'a set" where +definition\<^marker>\tag important\ space :: "'a measure \ 'a set" where "space M = fst (Rep_measure M)" -definition%important sets :: "'a measure \ 'a set set" where +definition\<^marker>\tag important\ sets :: "'a measure \ 'a set set" where "sets M = fst (snd (Rep_measure M))" -definition%important emeasure :: "'a measure \ 'a set \ ennreal" where +definition\<^marker>\tag important\ emeasure :: "'a measure \ 'a set \ ennreal" where "emeasure M = snd (snd (Rep_measure M))" -definition%important measure :: "'a measure \ 'a set \ real" where +definition\<^marker>\tag important\ measure :: "'a measure \ 'a set \ 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 \ 'a set set \ ('a set \ ennreal) \ 'a measure" +definition\<^marker>\tag important\ measure_of :: "'a set \ 'a set set \ ('a set \ ennreal) \ 'a measure" where "measure_of \ A \ = Abs_measure (\, if A \ Pow \ then sigma_sets \ A else {{}, \}, @@ -1740,7 +1740,7 @@ subsubsection \Measurable functions\ -definition%important measurable :: "'a measure \ 'b measure \ ('a \ 'b) set" +definition\<^marker>\tag important\ measurable :: "'a measure \ 'b measure \ ('a \ 'b) set" (infixr "\\<^sub>M" 60) where "measurable A B = {f \ space A \ space B. \y \ sets B. f -` y \ space A \ sets A}" @@ -1903,7 +1903,7 @@ subsubsection \Counting space\ -definition%important count_space :: "'a set \ 'a measure" where +definition\<^marker>\tag important\ count_space :: "'a set \ 'a measure" where "count_space \ = measure_of \ (Pow \) (\A. if finite A then of_nat (card A) else \)" lemma @@ -1979,7 +1979,7 @@ "space N = {} \ f \ measurable M N \ space M = {}" by (auto simp add: measurable_def Pi_iff) -subsubsection%unimportant \Extend measure\ +subsubsection\<^marker>\tag unimportant\ \Extend measure\ definition extend_measure :: "'a set \ 'b set \ ('b \ 'a set) \ ('b \ ennreal) \ 'a measure" where @@ -2043,7 +2043,7 @@ subsection \The smallest \\\-algebra regarding a function\ -definition%important vimage_algebra :: "'a set \ ('a \ 'b) \ 'b measure \ 'a measure" where +definition\<^marker>\tag important\ vimage_algebra :: "'a set \ ('a \ 'b) \ 'b measure \ 'a measure" where "vimage_algebra X f M = sigma X {f -` A \ X | A. A \ sets M}" lemma space_vimage_algebra[simp]: "space (vimage_algebra X f M) = X" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Starlike.thy --- 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 \Midpoint\ -definition%important midpoint :: "'a::real_vector \ 'a \ 'a" +definition\<^marker>\tag important\ midpoint :: "'a::real_vector \ 'a \ '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 \Line segments\ -definition%important closed_segment :: "'a::real_vector \ 'a \ 'a set" +definition\<^marker>\tag important\ closed_segment :: "'a::real_vector \ 'a \ 'a set" where "closed_segment a b = {(1 - u) *\<^sub>R a + u *\<^sub>R b | u::real. 0 \ u \ u \ 1}" -definition%important open_segment :: "'a::real_vector \ 'a \ 'a set" where +definition\<^marker>\tag important\ open_segment :: "'a::real_vector \ 'a \ 'a set" where "open_segment a b \ closed_segment a b - {a,b}" lemmas segment = open_segment_def closed_segment_def @@ -618,7 +618,7 @@ subsection\Starlike sets\ -definition%important "starlike S \ (\a\S. \x\S. closed_segment a x \ S)" +definition\<^marker>\tag important\ "starlike S \ (\a\S. \x\S. closed_segment a x \ 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\More results about segments\ +subsection\<^marker>\tag unimportant\\More results about segments\ lemma dist_half_times2: fixes a :: "'a :: real_normed_vector" @@ -903,7 +903,7 @@ subsection\Betweenness\ -definition%important "between = (\(a,b) x. x \ closed_segment a b)" +definition\<^marker>\tag important\ "between = (\(a,b) x. x \ closed_segment a b)" lemma betweenI: assumes "0 \ u" "u \ 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 \Shrinking towards the interior of a convex set\ +subsection\<^marker>\tag unimportant\ \Shrinking towards the interior of a convex set\ lemma mem_interior_convex_shrink: fixes S :: "'a::euclidean_space set" @@ -1247,7 +1247,7 @@ qed -subsection%unimportant \Some obvious but surprisingly hard simplex lemmas\ +subsection\<^marker>\tag unimportant\ \Some obvious but surprisingly hard simplex lemmas\ lemma simplex: assumes "finite S" @@ -1631,7 +1631,7 @@ qed -subsection%unimportant \Relative interior of convex set\ +subsection\<^marker>\tag unimportant\ \Relative interior of convex set\ lemma rel_interior_convex_nonempty_aux: fixes S :: "'n::euclidean_space set" @@ -2019,7 +2019,7 @@ subsection\The relative frontier of a set\ -definition%important "rel_frontier S = closure S - rel_interior S" +definition\<^marker>\tag important\ "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 \Relative interior and closure under common operations\ +subsubsection\<^marker>\tag unimportant\ \Relative interior and closure under common operations\ lemma rel_interior_inter_aux: "\{rel_interior S |S. S \ I} \ \I" proof - @@ -3085,7 +3085,7 @@ by (force simp: rel_frontier_def rel_interior_Times assms closure_Times) -subsubsection%unimportant \Relative interior of convex cone\ +subsubsection\<^marker>\tag unimportant\ \Relative interior of convex cone\ lemma cone_rel_interior: fixes S :: "'m::euclidean_space set" @@ -3358,7 +3358,7 @@ qed -subsection%unimportant \Convexity on direct sums\ +subsection\<^marker>\tag unimportant\ \Convexity on direct sums\ lemma closure_sum: fixes S T :: "'a::real_normed_vector set" @@ -3723,7 +3723,7 @@ using \y < x\ by (simp add: field_simps) qed simp -subsection%unimportant\Explicit formulas for interior and relative interior of convex hull\ +subsection\<^marker>\tag unimportant\\Explicit formulas for interior and relative interior of convex hull\ lemma at_within_cbox_finite: assumes "x \ box a b" "x \ 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\Similar results for closure and (relative or absolute) frontier\ +subsection\<^marker>\tag unimportant\\Similar results for closure and (relative or absolute) frontier\ lemma closure_convex_hull [simp]: fixes s :: "'a::euclidean_space set" @@ -4242,7 +4242,7 @@ subsection \Coplanarity, and collinearity in terms of affine hull\ -definition%important coplanar where +definition\<^marker>\tag important\ coplanar where "coplanar s \ \u v w. s \ affine hull {u,v,w}" lemma collinear_affine_hull: @@ -4693,7 +4693,7 @@ done -subsection%unimportant\Basic lemmas about hyperplanes and halfspaces\ +subsection\<^marker>\tag unimportant\\Basic lemmas about hyperplanes and halfspaces\ lemma halfspace_Int_eq: "{x. a \ x \ b} \ {x. b \ a \ x} = {x. a \ x = b}" @@ -4751,9 +4751,9 @@ using halfspace_eq_empty_le [of "-a" "-b"] by simp -subsection%unimportant\Use set distance for an easy proof of separation properties\ - -proposition%unimportant separation_closures: +subsection\<^marker>\tag unimportant\\Use set distance for an easy proof of separation properties\ + +proposition\<^marker>\tag unimportant\ separation_closures: fixes S :: "'a::euclidean_space set" assumes "S \ closure T = {}" "T \ closure S = {}" obtains U V where "U \ V = {}" "open U" "open V" "S \ U" "T \ V" @@ -5079,7 +5079,7 @@ by (simp add: continuous_on_closed * closedin_imp_subset) qed -subsection%unimportant\Trivial fact: convexity equals connectedness for collinear sets\ +subsection\<^marker>\tag unimportant\\Trivial fact: convexity equals connectedness for collinear sets\ lemma convex_connected_collinear: fixes S :: "'a::euclidean_space set" @@ -5297,9 +5297,9 @@ by (simp add: clo closedin_closed_Int) qed -subsubsection%unimportant\Representing affine hull as a finite intersection of hyperplanes\ - -proposition%unimportant affine_hull_convex_Int_nonempty_interior: +subsubsection\<^marker>\tag unimportant\\Representing affine hull as a finite intersection of hyperplanes\ + +proposition\<^marker>\tag unimportant\ affine_hull_convex_Int_nonempty_interior: fixes S :: "'a::real_normed_vector set" assumes "convex S" "S \ interior T \ {}" shows "affine hull (S \ T) = affine hull S" @@ -5524,7 +5524,7 @@ shows "a \ 0 \ aff_dim {x. a \ x = r} = DIM('a) - 1" by (metis aff_dim_eq_hyperplane affine_hull_eq affine_hyperplane) -subsection%unimportant\Some stepping theorems\ +subsection\<^marker>\tag unimportant\\Some stepping theorems\ lemma aff_dim_insert: fixes a :: "'a::euclidean_space" @@ -5649,9 +5649,9 @@ by (metis One_nat_def \a \ 0\ affine_bounded_eq_lowdim affine_hyperplane bounded_hyperplane_eq_trivial_0) qed -subsection%unimportant\General case without assuming closure and getting non-strict separation\ - -proposition%unimportant separating_hyperplane_closed_point_inset: +subsection\<^marker>\tag unimportant\\General case without assuming closure and getting non-strict separation\ + +proposition\<^marker>\tag unimportant\ separating_hyperplane_closed_point_inset: fixes S :: "'a::euclidean_space set" assumes "convex S" "closed S" "S \ {}" "z \ S" obtains a b where "a \ S" "(a - z) \ z < b" "\x. x \ S \ b < (a - z) \ x" @@ -5688,7 +5688,7 @@ by simp (metis \0 \ S\) -proposition%unimportant separating_hyperplane_set_0_inspan: +proposition\<^marker>\tag unimportant\ separating_hyperplane_set_0_inspan: fixes S :: "'a::euclidean_space set" assumes "convex S" "S \ {}" "0 \ S" obtains a where "a \ span S" "a \ 0" "\x. x \ S \ 0 \ a \ x" @@ -5779,7 +5779,7 @@ done qed -proposition%unimportant supporting_hyperplane_rel_boundary: +proposition\<^marker>\tag unimportant\ supporting_hyperplane_rel_boundary: fixes S :: "'a::euclidean_space set" assumes "convex S" "x \ S" and xno: "x \ rel_interior S" obtains a where "a \ 0" @@ -5837,7 +5837,7 @@ by (metis assms convex_closure convex_rel_interior_closure) -subsection%unimportant\ Some results on decomposing convex hulls: intersections, simplicial subdivision\ +subsection\<^marker>\tag unimportant\\ Some results on decomposing convex hulls: intersections, simplicial subdivision\ lemma fixes s :: "'a::euclidean_space set" @@ -5882,7 +5882,7 @@ qed -proposition%unimportant affine_hull_Int: +proposition\<^marker>\tag unimportant\ affine_hull_Int: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent(s \ t)" shows "affine hull (s \ t) = affine hull s \ 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>\tag unimportant\ convex_hull_Int: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent(s \ t)" shows "convex hull (s \ t) = convex hull s \ convex hull t" @@ -5898,7 +5898,7 @@ apply (simp add: hull_mono) by (simp add: convex_hull_Int_subset assms) -proposition%unimportant +proposition\<^marker>\tag unimportant\ fixes s :: "'a::euclidean_space set set" assumes "\ affine_dependent (\s)" shows affine_hull_Inter: "affine hull (\s) = (\t\s. affine hull t)" (is "?A") @@ -5928,7 +5928,7 @@ by auto qed -proposition%unimportant in_convex_hull_exchange_unique: +proposition\<^marker>\tag unimportant\ in_convex_hull_exchange_unique: fixes S :: "'a::euclidean_space set" assumes naff: "\ affine_dependent S" and a: "a \ convex hull S" and S: "T \ S" "T' \ S" @@ -6076,7 +6076,7 @@ qed qed -corollary%unimportant convex_hull_exchange_Int: +corollary\<^marker>\tag unimportant\ convex_hull_exchange_Int: fixes a :: "'a::euclidean_space" assumes "\ affine_dependent S" "a \ convex hull S" "T \ S" "T' \ S" shows "(convex hull (insert a T)) \ (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>\tag unimportant\ affine_hyperplane_sums_eq_UNIV: fixes S :: "'a :: euclidean_space set" assumes "affine S" and "S \ {v. a \ v = b} \ {}" @@ -6466,7 +6466,7 @@ by force qed -corollary%unimportant dense_complement_affine: +corollary\<^marker>\tag unimportant\ 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 \ T = {}") @@ -6487,7 +6487,7 @@ by (metis closure_translation translation_diff translation_invert) qed -corollary%unimportant dense_complement_openin_affine_hull: +corollary\<^marker>\tag unimportant\ 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>\tag unimportant\ 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 \convex S\ closure_mono convex_closure_rel_interior order_refl rel_interior_subset) qed -corollary%unimportant dense_complement_convex_closed: +corollary\<^marker>\tag unimportant\ 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\Parallel slices, etc\ +subsection\<^marker>\tag unimportant\\Parallel slices, etc\ text\ 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.\ -proposition%unimportant affine_parallel_slice: +proposition\<^marker>\tag unimportant\ affine_parallel_slice: fixes S :: "'a :: euclidean_space set" assumes "affine S" and "S \ {x. a \ x \ b} \ {}" @@ -6845,7 +6845,7 @@ qed qed -corollary%unimportant paracompact_closed: +corollary\<^marker>\tag unimportant\ paracompact_closed: fixes S :: "'a :: {metric_space,second_countable_topology} set" assumes "closed S" and opC: "\T. T \ \ \ open T" @@ -6857,7 +6857,7 @@ by (rule paracompact_closedin [of UNIV S \]) (auto simp: assms) -subsection%unimportant\Closed-graph characterization of continuity\ +subsection\<^marker>\tag unimportant\\Closed-graph characterization of continuity\ lemma continuous_closed_graph_gen: fixes T :: "'b::real_normed_vector set" @@ -6971,9 +6971,9 @@ by simp qed -subsection%unimportant\The union of two collinear segments is another segment\ - -proposition%unimportant in_convex_hull_exchange: +subsection\<^marker>\tag unimportant\\The union of two collinear segments is another segment\ + +proposition\<^marker>\tag unimportant\ in_convex_hull_exchange: fixes a :: "'a::euclidean_space" assumes a: "a \ convex hull S" and xS: "x \ convex hull S" obtains b where "b \ S" "x \ convex hull (insert a (S - {b}))" @@ -7250,7 +7250,7 @@ subsection\Orthogonal complement\ -definition%important orthogonal_comp ("_\<^sup>\" [80] 80) +definition\<^marker>\tag important\ orthogonal_comp ("_\<^sup>\" [80] 80) where "orthogonal_comp W \ {x. \y \ W. orthogonal y x}" proposition subspace_orthogonal_comp: "subspace (W\<^sup>\)" @@ -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 \A non-injective linear function maps into a hyperplane.\ +subsection\<^marker>\tag unimportant\ \A non-injective linear function maps into a hyperplane.\ lemma linear_surj_adj_imp_inj: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Summation_Tests.thy --- 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. \ -definition%important conv_radius :: "(nat \ 'a :: banach) \ ereal" where +definition\<^marker>\tag important\ conv_radius :: "(nat \ 'a :: banach) \ ereal" where "conv_radius f = inverse (limsup (\n. ereal (root n (norm (f n)))))" lemma conv_radius_cong_weak [cong]: "(\n. f n = g n) \ conv_radius f = conv_radius g" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Tagged_Division.thy --- 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 \Sundries\ +subsection\<^marker>\tag unimportant\ \Sundries\ text\A transitive relation is well-founded if all initial segments are finite.\ @@ -132,10 +132,10 @@ subsection \Bounds on intervals where they exist\ -definition%important interval_upperbound :: "('a::euclidean_space) set \ 'a" +definition\<^marker>\tag important\ interval_upperbound :: "('a::euclidean_space) set \ 'a" where "interval_upperbound s = (\i\Basis. (SUP x\s. x\i) *\<^sub>R i)" -definition%important interval_lowerbound :: "('a::euclidean_space) set \ 'a" +definition\<^marker>\tag important\ interval_lowerbound :: "('a::euclidean_space) set \ 'a" where "interval_lowerbound s = (\i\Basis. (INF x\s. x\i) *\<^sub>R i)" lemma interval_upperbound[simp]: @@ -194,7 +194,7 @@ subsection \The notion of a gauge --- simply an open set containing the point\ -definition%important "gauge \ \ (\x. x \ \ x \ open (\ x))" +definition\<^marker>\tag important\ "gauge \ \ (\x. x \ \ x \ open (\ x))" lemma gaugeI: assumes "\x. x \ \ x" @@ -251,7 +251,7 @@ subsection \Divisions\ -definition%important division_of (infixl "division'_of" 40) +definition\<^marker>\tag important\ division_of (infixl "division'_of" 40) where "s division_of i \ finite s \ @@ -997,7 +997,7 @@ subsection \Tagged (partial) divisions\ -definition%important tagged_partial_division_of (infixr "tagged'_partial'_division'_of" 40) +definition\<^marker>\tag important\ tagged_partial_division_of (infixr "tagged'_partial'_division'_of" 40) where "s tagged_partial_division_of i \ finite s \ (\x K. (x, K) \ s \ x \ K \ K \ i \ (\a b. K = cbox a b)) \ @@ -1014,7 +1014,7 @@ (x2, K2) \ s \ (x1, K1) \ (x2, K2) \ interior K1 \ interior K2 = {}" using assms unfolding tagged_partial_division_of_def by blast+ -definition%important tagged_division_of (infixr "tagged'_division'_of" 40) +definition\<^marker>\tag important\ tagged_division_of (infixr "tagged'_division'_of" 40) where "s tagged_division_of i \ s tagged_partial_division_of i \ (\{K. \x. (x,K) \ s} = i)" lemma tagged_division_of_finite: "s tagged_division_of i \ finite s" @@ -1292,9 +1292,9 @@ \operative_division\. Instances for the monoid are \<^typ>\'a option\, \<^typ>\real\, and \<^typ>\bool\.\ -paragraph%important \Using additivity of lifted function to encode definedness.\ +paragraph\<^marker>\tag important\ \Using additivity of lifted function to encode definedness.\ text \%whitespace\ -definition%important lift_option :: "('a \ 'b \ 'c) \ 'a option \ 'b option \ 'c option" +definition\<^marker>\tag important\ lift_option :: "('a \ 'b \ 'c) \ 'a option \ 'b option \ 'c option" where "lift_option f a' b' = Option.bind a' (\a. Option.bind b' (\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>\tag important\ comm_monoid_lift_option: assumes "comm_monoid f z" shows "comm_monoid (lift_option f) (Some z)" proof - @@ -1335,7 +1335,7 @@ paragraph \Division points\ text \%whitespace\ -definition%important "division_points (k::('a::euclidean_space) set) d = +definition\<^marker>\tag important\ "division_points (k::('a::euclidean_space) set) d = {(j,x). j \ Basis \ (interval_lowerbound k)\j < x \ x < (interval_upperbound k)\j \ (\i\d. (interval_lowerbound i)\j = x \ (interval_upperbound i)\j = x)}" @@ -1858,7 +1858,7 @@ subsection \Fine-ness of a partition w.r.t. a gauge\ -definition%important fine (infixr "fine" 46) +definition\<^marker>\tag important\ fine (infixr "fine" 46) where "d fine s \ (\(x,k) \ s. k \ d x)" lemma fineI: @@ -1907,7 +1907,7 @@ qed (* FIXME structure here, do not have one lemma in each subsection *) -subsection%unimportant \The set we're concerned with must be closed\ +subsection\<^marker>\tag unimportant\ \The set we're concerned with must be closed\ lemma division_of_closed: fixes i :: "'n::euclidean_space set" @@ -2301,7 +2301,7 @@ with ptag that show ?thesis by auto qed -subsubsection%important \Covering lemma\ +subsubsection\<^marker>\tag important\ \Covering lemma\ text\ 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 \Divisions over all gauges towards finer divisions.\ -definition%important division_filter :: "'a::euclidean_space set \ ('a \ 'a set) set filter" +definition\<^marker>\tag important\ division_filter :: "'a::euclidean_space set \ ('a \ 'a set) set filter" where "division_filter s = (INF g\{g. gauge g}. principal {p. p tagged_division_of s \ g fine p})" proposition eventually_division_filter: diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Topology_Euclidean_Space.thy --- 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\Balls in Euclidean Space\ +subsection\<^marker>\tag unimportant\\Balls in Euclidean Space\ lemma cball_subset_cball_iff: fixes a :: "'a :: euclidean_space" @@ -382,11 +382,11 @@ corollary Zero_neq_One[iff]: "0 \ One" by (metis One_non_0) -definition%important (in euclidean_space) eucl_less (infix "\tag important\ (in euclidean_space) eucl_less (infix " (\i\Basis. a \ i < b \ i)" -definition%important box_eucl_less: "box a b = {x. a x i\Basis. a \ i \ x \ i \ x \ i \ b \ i}" +definition\<^marker>\tag important\ box_eucl_less: "box a b = {x. a x \tag important\ "cbox a b = {x. \i\Basis. a \ i \ x \ i \ x \ i \ b \ i}" lemma box_def: "box a b = {x. \i\Basis. a \ i < x \ i \ x \ i < b \ i}" and in_box_eucl_less: "x \ box a b \ a x General Intervals\ -definition%important "is_interval (s::('a::euclidean_space) set) \ +definition\<^marker>\tag important\ "is_interval (s::('a::euclidean_space) set) \ (\a\s. \b\s. \x. (\i\Basis. ((a\i \ x\i \ x\i \ b\i) \ (b\i \ x\i \ x\i \ a\i))) \ x \ s)" lemma is_interval_1: @@ -1119,7 +1119,7 @@ by (metis image_cong uminus_add_conv_diff) -subsection%unimportant \Bounded Projections\ +subsection\<^marker>\tag unimportant\ \Bounded Projections\ 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 \Structural rules for pointwise continuity\ +subsection\<^marker>\tag unimportant\ \Structural rules for pointwise continuity\ lemma continuous_infnorm[continuous_intros]: "continuous F f \ continuous F (\x. infnorm (f x))" @@ -1145,7 +1145,7 @@ using assms unfolding continuous_def by (rule tendsto_inner) -subsection%unimportant \Structural rules for setwise continuity\ +subsection\<^marker>\tag unimportant\ \Structural rules for setwise continuity\ lemma continuous_on_infnorm[continuous_intros]: "continuous_on s f \ continuous_on s (\x. infnorm (f x))" @@ -1160,7 +1160,7 @@ by (rule bounded_bilinear.continuous_on) -subsection%unimportant \Openness of halfspaces.\ +subsection\<^marker>\tag unimportant\ \Openness of halfspaces.\ 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 Closure of halfspaces and hyperplanes\ +subsection\<^marker>\tag unimportant\ \Closure of halfspaces and hyperplanes\ 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\Some more convenient intermediate-value theorem formulations\ +subsection\<^marker>\tag unimportant\\Some more convenient intermediate-value theorem formulations\ lemma connected_ivt_hyperplane: assumes "connected S" and xy: "x \ S" "y \ S" and b: "inner a x \ b" "b \ inner a y" @@ -1537,8 +1537,8 @@ (auto intro!: assms bounded_linear_inner_left bounded_linear_image simp: euclidean_representation) -instance%important euclidean_space \ heine_borel -proof%unimportant +instance\<^marker>\tag important\ euclidean_space \ heine_borel +proof fix f :: "nat \ '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 \ banach .. +instance\<^marker>\tag important\ euclidean_space \ banach .. instance euclidean_space \ second_countable_topology proof @@ -1678,7 +1678,7 @@ qed -subsection%unimportant\Componentwise limits and continuity\ +subsection\<^marker>\tag unimportant\\Componentwise limits and continuity\ text\But is the premise really necessary? Need to generalise @{thm euclidean_dist_l2}\ lemma Euclidean_dist_upper: "i \ Basis \ dist (x \ i) (y \ i) \ dist x y" @@ -1796,7 +1796,7 @@ by (force simp: bounded_linear_def bounded_linear_axioms_def \linear f'\) qed -subsection%unimportant \Continuous Extension\ +subsection\<^marker>\tag unimportant\ \Continuous Extension\ definition clamp :: "'a::euclidean_space \ 'a \ 'a \ 'a" where "clamp a b x = (if (\i\Basis. a \ i \ b \ i) @@ -1985,7 +1985,7 @@ qed -subsection%unimportant \Diameter\ +subsection\<^marker>\tag unimportant\ \Diameter\ 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\Relating linear images to open/closed/interior/closure/connected\ +subsection\<^marker>\tag unimportant\\Relating linear images to open/closed/interior/closure/connected\ proposition open_surjective_linear_image: fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" @@ -2155,7 +2155,7 @@ using connected_continuous_image assms linear_continuous_on linear_conv_bounded_linear by blast -subsection%unimportant \"Isometry" (up to constant bounds) of Injective Linear Map\ +subsection\<^marker>\tag unimportant\ \"Isometry" (up to constant bounds) of Injective Linear Map\ proposition injective_imp_isometric: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" @@ -2243,7 +2243,7 @@ qed -subsection%unimportant \Some properties of a canonical subspace\ +subsection\<^marker>\tag unimportant\ \Some properties of a canonical subspace\ lemma closed_substandard: "closed {x::'a::euclidean_space. \i\Basis. P i \ x\i = 0}" (is "closed ?A") diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Uniform_Limit.thy --- 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 \Definition\ -definition%important uniformly_on :: "'a set \ ('a \ 'b::metric_space) \ ('a \ 'b) filter" +definition\<^marker>\tag important\ uniformly_on :: "'a set \ ('a \ 'b::metric_space) \ ('a \ 'b) filter" where "uniformly_on S l = (INF e\{0 <..}. principal {f. \x\S. dist (f x) (l x) < e})" -abbreviation%important +abbreviation\<^marker>\tag important\ "uniform_limit S f l \ filterlim f (uniformly_on S l)" definition uniformly_convergent_on where @@ -366,7 +366,7 @@ qed text\Alternative version, formulated as in HOL Light\ -corollary%unimportant series_comparison_uniform: +corollary\<^marker>\tag unimportant\ series_comparison_uniform: fixes f :: "_ \ nat \ _ :: banach" assumes g: "summable g" and le: "\n x. N \ n \ x \ A \ norm(f x n) \ g n" shows "\l. \e. 0 < e \ (\N. \n x. N \ n \ x \ A \ dist(sum (f x) {..\tag unimportant\ Weierstrass_m_test: fixes f :: "_ \ _ \ _ :: banach" assumes "\n x. x \ A \ norm (f n x) \ M n" assumes "summable M" shows "uniform_limit A (\n x. \ix. suminf (\i. f i x)) sequentially" using assms by (intro Weierstrass_m_test_ev always_eventually) auto -corollary%unimportant Weierstrass_m_test'_ev: +corollary\<^marker>\tag unimportant\ Weierstrass_m_test'_ev: fixes f :: "_ \ _ \ _ :: banach" assumes "eventually (\n. \x\A. norm (f n x) \ M n) sequentially" "summable M" shows "uniformly_convergent_on A (\n x. \i\tag unimportant\ Weierstrass_m_test': fixes f :: "_ \ _ \ _ :: banach" assumes "\n x. x \ A \ norm (f n x) \ M n" "summable M" shows "uniformly_convergent_on A (\n x. \iStructural introduction rules\ +subsection\<^marker>\tag unimportant\ \Structural introduction rules\ named_theorems uniform_limit_intros "introduction rules for uniform_limit" setup \ diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Weierstrass_Theorems.thy --- 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 \Bernstein polynomials\ -definition%important Bernstein :: "[nat,nat,real] \ real" where +definition\<^marker>\tag important\ Bernstein :: "[nat,nat,real] \ real" where "Bernstein n k x \ of_nat (n choose k) * x ^ k * (1 - x) ^ (n - k)" lemma Bernstein_nonneg: "\0 \ x; x \ 1\ \ 0 \ Bernstein n k x" @@ -201,7 +201,7 @@ lemma prod: "\finite I; \i. i \ I \ f i \ R\ \ (\x. \i \ I. f i x) \ R" by (induct I rule: finite_induct; simp add: const mult) - definition%important normf :: "('a::t2_space \ real) \ real" + definition\<^marker>\tag important\ normf :: "('a::t2_space \ real) \ real" where "normf f \ SUP x\S. \f x\" lemma normf_upper: "\continuous_on S f; x \ S\ \ \f x\ \ normf f" @@ -792,7 +792,7 @@ declare real_polynomial_function.intros [intro] -definition%important polynomial_function :: "('a::real_normed_vector \ 'b::real_normed_vector) \ bool" +definition\<^marker>\tag important\ polynomial_function :: "('a::real_normed_vector \ 'b::real_normed_vector) \ bool" where "polynomial_function p \ (\f. bounded_linear f \ real_polynomial_function (f o p))" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Analysis/Winding_Numbers.thy --- 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>\tag important\ 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\A per-function version for continuous logs, a kind of monodromy\ -proposition%unimportant winding_number_compose_exp: +proposition\<^marker>\tag unimportant\ winding_number_compose_exp: assumes "path p" shows "winding_number (exp \ p) 0 = (pathfinish p - pathstart p) / (2 * of_real pi * \)" proof - @@ -1020,7 +1020,7 @@ finally show ?thesis . qed -subsection%unimportant \The winding number defines a continuous logarithm for the path itself\ +subsection\<^marker>\tag unimportant\ \The winding number defines a continuous logarithm for the path itself\ lemma winding_number_as_continuous_log: assumes "path p" and \: "\ \ path_image p" diff -r ad6d4a14adb5 -r f03a01a18c6e src/HOL/Homology/Invariance_of_Domain.thy --- 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 \ {}" shows "DIM('a) \ 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 \ 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 \ 'b::euclidean_space" @@ -2438,8 +2438,8 @@ fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "DIM('b) \ 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 \ 'b::euclidean_space"