src/HOL/Analysis/Change_Of_Vars.thy
changeset 69720 be6634e99e09
parent 69683 8b3458ca0762
child 69739 8b47c021666e
--- a/src/HOL/Analysis/Change_Of_Vars.thy	Fri Dec 14 14:33:26 2018 +0100
+++ b/src/HOL/Analysis/Change_Of_Vars.thy	Tue Jan 22 21:13:23 2019 +0000
@@ -9,9 +9,9 @@
 
 begin
 
-subsection%unimportant \<open>Orthogonal Transformation of Balls\<close>
+subsection \<open>Orthogonal Transformation of Balls\<close>
 
-lemma%unimportant image_orthogonal_transformation_ball:
+lemma image_orthogonal_transformation_ball:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
   assumes "orthogonal_transformation f"
   shows "f ` ball x r = ball (f x) r"
@@ -27,7 +27,7 @@
     by (auto simp: orthogonal_transformation_isometry)
 qed
 
-lemma%unimportant  image_orthogonal_transformation_cball:
+lemma  image_orthogonal_transformation_cball:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
   assumes "orthogonal_transformation f"
   shows "f ` cball x r = cball (f x) r"
@@ -46,14 +46,14 @@
 
 subsection \<open>Measurable Shear and Stretch\<close>
 
-proposition%important
+proposition
   fixes a :: "real^'n"
   assumes "m \<noteq> n" and ab_ne: "cbox a b \<noteq> {}" and an: "0 \<le> a$n"
   shows measurable_shear_interval: "(\<lambda>x. \<chi> i. if i = m then x$m + x$n else x$i) ` (cbox a b) \<in> lmeasurable"
        (is  "?f ` _ \<in> _")
    and measure_shear_interval: "measure lebesgue ((\<lambda>x. \<chi> i. if i = m then x$m + x$n else x$i) ` cbox a b)
                = measure lebesgue (cbox a b)" (is "?Q")
-proof%unimportant -
+proof -
   have lin: "linear ?f"
     by (rule linearI) (auto simp: plus_vec_def scaleR_vec_def algebra_simps)
   show fab: "?f ` cbox a b \<in> lmeasurable"
@@ -155,13 +155,13 @@
 qed
 
 
-proposition%important
+proposition
   fixes S :: "(real^'n) set"
   assumes "S \<in> lmeasurable"
   shows measurable_stretch: "((\<lambda>x. \<chi> k. m k * x$k) ` S) \<in> lmeasurable" (is  "?f ` S \<in> _")
     and measure_stretch: "measure lebesgue ((\<lambda>x. \<chi> k. m k * x$k) ` S) = \<bar>prod m UNIV\<bar> * measure lebesgue S"
     (is "?MEQ")
-proof%unimportant -
+proof -
   have "(?f ` S) \<in> lmeasurable \<and> ?MEQ"
   proof (cases "\<forall>k. m k \<noteq> 0")
     case True
@@ -259,12 +259,12 @@
 qed
 
 
-proposition%important
+proposition
  fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
   assumes "linear f" "S \<in> lmeasurable"
   shows measurable_linear_image: "(f ` S) \<in> lmeasurable"
     and measure_linear_image: "measure lebesgue (f ` S) = \<bar>det (matrix f)\<bar> * measure lebesgue S" (is "?Q f S")
-proof%unimportant -
+proof -
   have "\<forall>S \<in> lmeasurable. (f ` S) \<in> lmeasurable \<and> ?Q f S"
   proof (rule induct_linear_elementary [OF \<open>linear f\<close>]; intro ballI)
     fix f g and S :: "(real,'n) vec set"
@@ -407,7 +407,7 @@
 qed
 
 
-lemma%unimportant
+lemma
  fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
   assumes f: "orthogonal_transformation f" and S: "S \<in> lmeasurable"
   shows measurable_orthogonal_image: "f ` S \<in> lmeasurable"
@@ -430,10 +430,10 @@
 inductive%important gdelta :: "'a::topological_space set \<Rightarrow> bool" where
   "(\<And>n::nat. open (F n)) \<Longrightarrow> gdelta (\<Inter>(F ` UNIV))"
 
-lemma%important fsigma_Union_compact:
+lemma fsigma_Union_compact:
   fixes S :: "'a::{real_normed_vector,heine_borel} set"
   shows "fsigma S \<longleftrightarrow> (\<exists>F::nat \<Rightarrow> 'a set. range F \<subseteq> Collect compact \<and> S = \<Union>(F ` UNIV))"
-proof%unimportant safe
+proof safe
   assume "fsigma S"
   then obtain F :: "nat \<Rightarrow> 'a set" where F: "range F \<subseteq> Collect closed" "S = \<Union>(F ` UNIV)"
     by (meson fsigma.cases image_subsetI mem_Collect_eq)
@@ -464,7 +464,7 @@
     by (simp add: compact_imp_closed fsigma.intros image_subset_iff)
 qed
 
-lemma%unimportant gdelta_imp_fsigma: "gdelta S \<Longrightarrow> fsigma (- S)"
+lemma gdelta_imp_fsigma: "gdelta S \<Longrightarrow> fsigma (- S)"
 proof (induction rule: gdelta.induct)
   case (1 F)
   have "- \<Inter>(F ` UNIV) = (\<Union>i. -(F i))"
@@ -473,7 +473,7 @@
     by (simp add: fsigma.intros closed_Compl 1)
 qed
 
-lemma%unimportant fsigma_imp_gdelta: "fsigma S \<Longrightarrow> gdelta (- S)"
+lemma fsigma_imp_gdelta: "fsigma S \<Longrightarrow> gdelta (- S)"
 proof (induction rule: fsigma.induct)
   case (1 F)
   have "- \<Union>(F ` UNIV) = (\<Inter>i. -(F i))"
@@ -482,11 +482,11 @@
     by (simp add: 1 gdelta.intros open_closed)
 qed
 
-lemma%unimportant gdelta_complement: "gdelta(- S) \<longleftrightarrow> fsigma S"
+lemma gdelta_complement: "gdelta(- S) \<longleftrightarrow> fsigma S"
   using fsigma_imp_gdelta gdelta_imp_fsigma by force
 
 text\<open>A Lebesgue set is almost an \<open>F_sigma\<close> or \<open>G_delta\<close>.\<close>
-lemma%unimportant lebesgue_set_almost_fsigma:
+lemma lebesgue_set_almost_fsigma:
   assumes "S \<in> sets lebesgue"
   obtains C T where "fsigma C" "negligible T" "C \<union> T = S" "disjnt C T"
 proof -
@@ -521,7 +521,7 @@
   qed
 qed
 
-lemma%unimportant lebesgue_set_almost_gdelta:
+lemma lebesgue_set_almost_gdelta:
   assumes "S \<in> sets lebesgue"
   obtains C T where "gdelta C" "negligible T" "S \<union> T = C" "disjnt S T"
 proof -
@@ -539,12 +539,12 @@
 qed
 
 
-proposition%important measure_semicontinuous_with_hausdist_explicit:
+proposition measure_semicontinuous_with_hausdist_explicit:
   assumes "bounded S" and neg: "negligible(frontier S)" and "e > 0"
   obtains d where "d > 0"
                   "\<And>T. \<lbrakk>T \<in> lmeasurable; \<And>y. y \<in> T \<Longrightarrow> \<exists>x. x \<in> S \<and> dist x y < d\<rbrakk>
                         \<Longrightarrow> measure lebesgue T < measure lebesgue S + e"
-proof%unimportant (cases "S = {}")
+proof (cases "S = {}")
   case True
   with that \<open>e > 0\<close> show ?thesis by force
 next
@@ -608,10 +608,10 @@
   qed
 qed
 
-proposition%important lebesgue_regular_inner:
+proposition lebesgue_regular_inner:
  assumes "S \<in> sets lebesgue"
  obtains K C where "negligible K" "\<And>n::nat. compact(C n)" "S = (\<Union>n. C n) \<union> K"
-proof%unimportant -
+proof -
   have "\<exists>T. closed T \<and> T \<subseteq> S \<and> (S - T) \<in> lmeasurable \<and> measure lebesgue (S - T) < (1/2)^n" for n
     using sets_lebesgue_inner_closed assms
     by (metis sets_lebesgue_inner_closed zero_less_divide_1_iff zero_less_numeral zero_less_power)
@@ -661,7 +661,7 @@
 qed
 
 
-lemma%unimportant sets_lebesgue_continuous_image:
+lemma sets_lebesgue_continuous_image:
   assumes T: "T \<in> sets lebesgue" and contf: "continuous_on S f"
     and negim: "\<And>T. \<lbrakk>negligible T; T \<subseteq> S\<rbrakk> \<Longrightarrow> negligible(f ` T)" and "T \<subseteq> S"
  shows "f ` T \<in> sets lebesgue"
@@ -686,7 +686,7 @@
     by (simp add: Teq image_Un image_Union)
 qed
 
-lemma%unimportant differentiable_image_in_sets_lebesgue:
+lemma differentiable_image_in_sets_lebesgue:
   fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
   assumes S: "S \<in> sets lebesgue" and dim: "DIM('m) \<le> DIM('n)" and f: "f differentiable_on S"
   shows "f`S \<in> sets lebesgue"
@@ -698,7 +698,7 @@
     by (auto simp: intro!: negligible_differentiable_image_negligible [OF dim])
 qed auto
 
-lemma%unimportant sets_lebesgue_on_continuous_image:
+lemma sets_lebesgue_on_continuous_image:
   assumes S: "S \<in> sets lebesgue" and X: "X \<in> sets (lebesgue_on S)" and contf: "continuous_on S f"
     and negim: "\<And>T. \<lbrakk>negligible T; T \<subseteq> S\<rbrakk> \<Longrightarrow> negligible(f ` T)"
   shows "f ` X \<in> sets (lebesgue_on (f ` S))"
@@ -713,7 +713,7 @@
     by (auto simp: sets_restrict_space_iff)
 qed
 
-lemma%unimportant differentiable_image_in_sets_lebesgue_on:
+lemma differentiable_image_in_sets_lebesgue_on:
   fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
   assumes S: "S \<in> sets lebesgue" and X: "X \<in> sets (lebesgue_on S)" and dim: "DIM('m) \<le> DIM('n)"
        and f: "f differentiable_on S"
@@ -727,7 +727,7 @@
 qed
 
 
-proposition%important
+proposition
  fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
   assumes S: "S \<in> lmeasurable"
   and deriv: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
@@ -737,7 +737,7 @@
        "f ` S \<in> lmeasurable"
     and measure_bounded_differentiable_image:
        "measure lebesgue (f ` S) \<le> B * measure lebesgue S" (is "?M")
-proof%unimportant -
+proof -
   have "f ` S \<in> lmeasurable \<and> measure lebesgue (f ` S) \<le> B * measure lebesgue S"
   proof (cases "B < 0")
     case True
@@ -963,13 +963,13 @@
   then show "f ` S \<in> lmeasurable" ?M by blast+
 qed
 
-lemma%important
+lemma (*FIXME needs name? *)
  fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
   assumes S: "S \<in> lmeasurable"
     and deriv: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
     and int: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on S"
   shows m_diff_image_weak: "f ` S \<in> lmeasurable \<and> measure lebesgue (f ` S) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
-proof%unimportant -
+proof -
   let ?\<mu> = "measure lebesgue"
   have aint_S: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) absolutely_integrable_on S"
     using int unfolding absolutely_integrable_on_def by auto
@@ -1161,7 +1161,7 @@
 qed
 
 
-theorem%important
+theorem(* FIXME needs name? *)
  fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
   assumes S: "S \<in> sets lebesgue"
     and deriv: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
@@ -1169,7 +1169,7 @@
   shows measurable_differentiable_image: "f ` S \<in> lmeasurable"
     and measure_differentiable_image:
        "measure lebesgue (f ` S) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)" (is "?M")
-proof%unimportant -
+proof -
   let ?I = "\<lambda>n::nat. cbox (vec (-n)) (vec n) \<inter> S"
   let ?\<mu> = "measure lebesgue"
   have "x \<in> cbox (vec (- real (nat \<lceil>norm x\<rceil>))) (vec (real (nat \<lceil>norm x\<rceil>)))" for x :: "real^'n::_"
@@ -1214,7 +1214,7 @@
 qed
 
 
-lemma%unimportant borel_measurable_simple_function_limit_increasing:
+lemma borel_measurable_simple_function_limit_increasing:
   fixes f :: "'a::euclidean_space \<Rightarrow> real"
   shows "(f \<in> borel_measurable lebesgue \<and> (\<forall>x. 0 \<le> f x)) \<longleftrightarrow>
          (\<exists>g. (\<forall>n x. 0 \<le> g n x \<and> g n x \<le> f x) \<and> (\<forall>n x. g n x \<le> (g(Suc n) x)) \<and>
@@ -1413,7 +1413,7 @@
 
 subsection\<open>Borel measurable Jacobian determinant\<close>
 
-lemma%unimportant lemma_partial_derivatives0:
+lemma lemma_partial_derivatives0:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "linear f" and lim0: "((\<lambda>x. f x /\<^sub>R norm x) \<longlongrightarrow> 0) (at 0 within S)"
     and lb: "\<And>v. v \<noteq> 0 \<Longrightarrow> (\<exists>k>0. \<forall>e>0. \<exists>x. x \<in> S - {0} \<and> norm x < e \<and> k * norm x \<le> \<bar>v \<bullet> x\<bar>)"
@@ -1486,7 +1486,7 @@
         mem_Collect_eq module_hom_zero span_base span_raw_def)
 qed
 
-lemma%unimportant lemma_partial_derivatives:
+lemma lemma_partial_derivatives:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "linear f" and lim: "((\<lambda>x. f (x - a) /\<^sub>R norm (x - a)) \<longlongrightarrow> 0) (at a within S)"
     and lb: "\<And>v. v \<noteq> 0 \<Longrightarrow> (\<exists>k>0.  \<forall>e>0. \<exists>x \<in> S - {a}. norm(a - x) < e \<and> k * norm(a - x) \<le> \<bar>v \<bullet> (x - a)\<bar>)"
@@ -1504,12 +1504,12 @@
 qed
 
 
-proposition%important borel_measurable_partial_derivatives:
+proposition borel_measurable_partial_derivatives:
   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n"
   assumes S: "S \<in> sets lebesgue"
     and f: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
   shows "(\<lambda>x. (matrix(f' x)$m$n)) \<in> borel_measurable (lebesgue_on S)"
-proof%unimportant -
+proof -
   have contf: "continuous_on S f"
     using continuous_on_eq_continuous_within f has_derivative_continuous by blast
   have "{x \<in> S.  (matrix (f' x)$m$n) \<le> b} \<in> sets lebesgue" for b
@@ -2095,7 +2095,7 @@
 qed
 
 
-theorem%important borel_measurable_det_Jacobian:
+theorem borel_measurable_det_Jacobian:
  fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
   assumes S: "S \<in> sets lebesgue" and f: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
   shows "(\<lambda>x. det(matrix(f' x))) \<in> borel_measurable (lebesgue_on S)"
@@ -2104,12 +2104,12 @@
 
 text\<open>The localisation wrt S uses the same argument for many similar results.\<close>
 (*See HOL Light's MEASURABLE_ON_LEBESGUE_MEASURABLE_PREIMAGE_BOREL, etc.*)
-lemma%important borel_measurable_lebesgue_on_preimage_borel:
+theorem borel_measurable_lebesgue_on_preimage_borel:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "S \<in> sets lebesgue"
   shows "f \<in> borel_measurable (lebesgue_on S) \<longleftrightarrow>
          (\<forall>T. T \<in> sets borel \<longrightarrow> {x \<in> S. f x \<in> T} \<in> sets lebesgue)"
-proof%unimportant -
+proof -
   have "{x. (if x \<in> S then f x else 0) \<in> T} \<in> sets lebesgue \<longleftrightarrow> {x \<in> S. f x \<in> T} \<in> sets lebesgue"
          if "T \<in> sets borel" for T
     proof (cases "0 \<in> T")
@@ -2131,7 +2131,7 @@
       by blast
 qed
 
-lemma%unimportant sets_lebesgue_almost_borel:
+lemma sets_lebesgue_almost_borel:
   assumes "S \<in> sets lebesgue"
   obtains B N where "B \<in> sets borel" "negligible N" "B \<union> N = S"
 proof -
@@ -2141,7 +2141,7 @@
     by (metis negligible_iff_null_sets negligible_subset null_sets_completionI that)
 qed
 
-lemma%unimportant double_lebesgue_sets:
+lemma double_lebesgue_sets:
  assumes S: "S \<in> sets lebesgue" and T: "T \<in> sets lebesgue" and fim: "f ` S \<subseteq> T"
  shows "(\<forall>U. U \<in> sets lebesgue \<and> U \<subseteq> T \<longrightarrow> {x \<in> S. f x \<in> U} \<in> sets lebesgue) \<longleftrightarrow>
           f \<in> borel_measurable (lebesgue_on S) \<and>
@@ -2186,7 +2186,7 @@
 
 subsection\<open>Simplest case of Sard's theorem (we don't need continuity of derivative)\<close>
 
-lemma%important Sard_lemma00:
+lemma Sard_lemma00:
   fixes P :: "'b::euclidean_space set"
   assumes "a \<ge> 0" and a: "a *\<^sub>R i \<noteq> 0" and i: "i \<in> Basis"
     and P: "P \<subseteq> {x. a *\<^sub>R i \<bullet> x = 0}"
@@ -2194,7 +2194,7 @@
  obtains S where "S \<in> lmeasurable"
             and "{z. norm z \<le> m \<and> (\<exists>t \<in> P. norm(z - t) \<le> e)} \<subseteq> S"
             and "measure lebesgue S \<le> (2 * e) * (2 * m) ^ (DIM('b) - 1)"
-proof%unimportant -
+proof -
   have "a > 0"
     using assms by simp
   let ?v = "(\<Sum>j\<in>Basis. (if j = i then e else m) *\<^sub>R j)"
@@ -2222,7 +2222,7 @@
 qed
 
 text\<open>As above, but reorienting the vector (HOL Light's @text{GEOM\_BASIS\_MULTIPLE\_TAC})\<close>
-lemma%unimportant Sard_lemma0:
+lemma Sard_lemma0:
   fixes P :: "(real^'n::{finite,wellorder}) set"
   assumes "a \<noteq> 0"
     and P: "P \<subseteq> {x. a \<bullet> x = 0}" and "0 \<le> m" "0 \<le> e"
@@ -2282,13 +2282,13 @@
 qed
 
 text\<open>As above, but translating the sets (HOL Light's @text{GEN\_GEOM\_ORIGIN\_TAC})\<close>
-lemma%important Sard_lemma1:
+lemma Sard_lemma1:
   fixes P :: "(real^'n::{finite,wellorder}) set"
    assumes P: "dim P < CARD('n)" and "0 \<le> m" "0 \<le> e"
  obtains S where "S \<in> lmeasurable"
             and "{z. norm(z - w) \<le> m \<and> (\<exists>t \<in> P. norm(z - w - t) \<le> e)} \<subseteq> S"
             and "measure lebesgue S \<le> (2 * e) * (2 * m) ^ (CARD('n) - 1)"
-proof%unimportant -
+proof -
   obtain a where "a \<noteq> 0" "P \<subseteq> {x. a \<bullet> x = 0}"
     using lowdim_subset_hyperplane [of P] P span_base by auto
   then obtain S where S: "S \<in> lmeasurable"
@@ -2306,7 +2306,7 @@
   qed
 qed
 
-lemma%important Sard_lemma2:
+lemma Sard_lemma2:
   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n::{finite,wellorder}"
   assumes mlen: "CARD('m) \<le> CARD('n)" (is "?m \<le> ?n")
     and "B > 0" "bounded S"
@@ -2314,7 +2314,7 @@
     and rank: "\<And>x. x \<in> S \<Longrightarrow> rank(matrix(f' x)) < CARD('n)"
     and B: "\<And>x. x \<in> S \<Longrightarrow> onorm(f' x) \<le> B"
   shows "negligible(f ` S)"
-proof%unimportant -
+proof -
   have lin_f': "\<And>x. x \<in> S \<Longrightarrow> linear(f' x)"
     using derS has_derivative_linear by blast
   show ?thesis
@@ -2521,13 +2521,13 @@
 qed
 
 
-theorem%important baby_Sard:
+theorem baby_Sard:
   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n::{finite,wellorder}"
   assumes mlen: "CARD('m) \<le> CARD('n)"
     and der: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
     and rank: "\<And>x. x \<in> S \<Longrightarrow> rank(matrix(f' x)) < CARD('n)"
   shows "negligible(f ` S)"
-proof%unimportant -
+proof -
   let ?U = "\<lambda>n. {x \<in> S. norm(x) \<le> n \<and> onorm(f' x) \<le> real n}"
   have "\<And>x. x \<in> S \<Longrightarrow> \<exists>n. norm x \<le> real n \<and> onorm (f' x) \<le> real n"
     by (meson linear order_trans real_arch_simple)
@@ -2549,7 +2549,7 @@
 
 subsection\<open>A one-way version of change-of-variables not assuming injectivity. \<close>
 
-lemma%important integral_on_image_ubound_weak:
+lemma integral_on_image_ubound_weak:
   fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real"
   assumes S: "S \<in> sets lebesgue"
       and f: "f \<in> borel_measurable (lebesgue_on (g ` S))"
@@ -2560,7 +2560,7 @@
     shows "f integrable_on (g ` S) \<and>
            integral (g ` S) f \<le> integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x))"
          (is "_ \<and> _ \<le> ?b")
-proof%unimportant -
+proof -
   let ?D = "\<lambda>x. \<bar>det (matrix (g' x))\<bar>"
   have cont_g: "continuous_on S g"
     using der_g has_derivative_continuous_on by blast
@@ -2738,14 +2738,14 @@
 qed
 
 
-lemma%important integral_on_image_ubound_nonneg:
+lemma integral_on_image_ubound_nonneg:
   fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real"
   assumes nonneg_fg: "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f(g x)"
       and der_g:   "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
       and intS: "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) integrable_on S"
   shows "f integrable_on (g ` S) \<and> integral (g ` S) f \<le> integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x))"
          (is "_ \<and> _ \<le> ?b")
-proof%unimportant -
+proof -
   let ?D = "\<lambda>x. det (matrix (g' x))"
   define S' where "S' \<equiv> {x \<in> S. ?D x * f(g x) \<noteq> 0}"
   then have der_gS': "\<And>x. x \<in> S' \<Longrightarrow> (g has_derivative g' x) (at x within S')"
@@ -2868,7 +2868,7 @@
 qed
 
 
-lemma%unimportant absolutely_integrable_on_image_real:
+lemma absolutely_integrable_on_image_real:
   fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
   assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
     and intS: "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) absolutely_integrable_on S"
@@ -2943,7 +2943,7 @@
 qed
 
 
-proposition%important absolutely_integrable_on_image:
+proposition absolutely_integrable_on_image:
   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
   assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
     and intS: "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S"
@@ -2951,7 +2951,7 @@
   apply (rule absolutely_integrable_componentwise [OF absolutely_integrable_on_image_real [OF der_g]])
   using%unimportant absolutely_integrable_component [OF intS]  by%unimportant auto
 
-proposition%important integral_on_image_ubound:
+proposition integral_on_image_ubound:
   fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
   assumes"\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f(g x)"
     and "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
@@ -2965,7 +2965,7 @@
 text\<open>The classic change-of-variables theorem. We have two versions with quite general hypotheses,
 the first that the transforming function has a continuous inverse, the second that the base set is
 Lebesgue measurable.\<close>
-lemma%unimportant cov_invertible_nonneg_le:
+lemma cov_invertible_nonneg_le:
   fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
   assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
     and der_h: "\<And>y. y \<in> T \<Longrightarrow> (h has_derivative h' y) (at y within T)"
@@ -3036,7 +3036,7 @@
 qed
 
 
-lemma%unimportant cov_invertible_nonneg_eq:
+lemma cov_invertible_nonneg_eq:
   fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
   assumes "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
       and "\<And>y. y \<in> T \<Longrightarrow> (h has_derivative h' y) (at y within T)"
@@ -3049,7 +3049,7 @@
   by (simp add: has_integral_iff) (meson eq_iff)
 
 
-lemma%unimportant cov_invertible_real:
+lemma cov_invertible_real:
   fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
   assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
       and der_h: "\<And>y. y \<in> T \<Longrightarrow> (h has_derivative h' y) (at y within T)"
@@ -3215,7 +3215,7 @@
 qed
 
 
-lemma%unimportant cv_inv_version3:
+lemma cv_inv_version3:
   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
   assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
     and der_h: "\<And>y. y \<in> T \<Longrightarrow> (h has_derivative h' y) (at y within T)"
@@ -3241,7 +3241,7 @@
 qed
 
 
-lemma%unimportant cv_inv_version4:
+lemma cv_inv_version4:
   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
   assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S) \<and> invertible(matrix(g' x))"
     and hg: "\<And>x. x \<in> S \<Longrightarrow> continuous_on (g ` S) h \<and> h(g x) = x"
@@ -3266,7 +3266,7 @@
 qed
 
 
-proposition%important has_absolute_integral_change_of_variables_invertible:
+theorem has_absolute_integral_change_of_variables_invertible:
   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
   assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
       and hg: "\<And>x. x \<in> S \<Longrightarrow> h(g x) = x"
@@ -3274,7 +3274,7 @@
   shows "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and> integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) = b \<longleftrightarrow>
          f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b"
     (is "?lhs = ?rhs")
-proof%unimportant -
+proof -
   let ?S = "{x \<in> S. invertible (matrix (g' x))}" and ?D = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)"
   have *: "?D absolutely_integrable_on ?S \<and> integral ?S ?D = b
            \<longleftrightarrow> f absolutely_integrable_on (g ` ?S) \<and> integral (g ` ?S) f = b"
@@ -3310,7 +3310,7 @@
 
 
 
-lemma%unimportant has_absolute_integral_change_of_variables_compact:
+theorem has_absolute_integral_change_of_variables_compact:
   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
   assumes "compact S"
       and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
@@ -3328,7 +3328,7 @@
 qed
 
 
-lemma%unimportant has_absolute_integral_change_of_variables_compact_family:
+lemma has_absolute_integral_change_of_variables_compact_family:
   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
   assumes compact: "\<And>n::nat. compact (F n)"
       and der_g: "\<And>x. x \<in> (\<Union>n. F n) \<Longrightarrow> (g has_derivative g' x) (at x within (\<Union>n. F n))"
@@ -3507,7 +3507,7 @@
 qed
 
 
-proposition%important has_absolute_integral_change_of_variables:
+theorem has_absolute_integral_change_of_variables:
   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
   assumes S: "S \<in> sets lebesgue"
     and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
@@ -3515,7 +3515,7 @@
   shows "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and>
            integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) = b
      \<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b"
-proof%unimportant -
+proof -
   obtain C N where "fsigma C" "negligible N" and CNS: "C \<union> N = S" and "disjnt C N"
     using lebesgue_set_almost_fsigma [OF S] .
   then obtain F :: "nat \<Rightarrow> (real^'m::_) set"
@@ -3569,7 +3569,7 @@
 qed
 
 
-corollary%important absolutely_integrable_change_of_variables:
+corollary absolutely_integrable_change_of_variables:
   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
   assumes "S \<in> sets lebesgue"
     and "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
@@ -3578,7 +3578,7 @@
      \<longleftrightarrow> (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S"
   using%unimportant assms has_absolute_integral_change_of_variables by%unimportant blast
 
-corollary%important integral_change_of_variables:
+corollary integral_change_of_variables:
   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
   assumes S: "S \<in> sets lebesgue"
     and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
@@ -3589,7 +3589,7 @@
   using%unimportant has_absolute_integral_change_of_variables [OF S der_g inj] disj
   by%unimportant blast
 
-lemma%unimportant has_absolute_integral_change_of_variables_1:
+lemma has_absolute_integral_change_of_variables_1:
   fixes f :: "real \<Rightarrow> real^'n::{finite,wellorder}" and g :: "real \<Rightarrow> real"
   assumes S: "S \<in> sets lebesgue"
     and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_vector_derivative g' x) (at x within S)"
@@ -3628,7 +3628,7 @@
 qed
 
 
-corollary%important absolutely_integrable_change_of_variables_1:
+corollary absolutely_integrable_change_of_variables_1:
   fixes f :: "real \<Rightarrow> real^'n::{finite,wellorder}" and g :: "real \<Rightarrow> real"
   assumes S: "S \<in> sets lebesgue"
     and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_vector_derivative g' x) (at x within S)"
@@ -3640,7 +3640,7 @@
 
 subsection\<open>Change of variables for integrals: special case of linear function\<close>
 
-lemma%unimportant has_absolute_integral_change_of_variables_linear:
+lemma has_absolute_integral_change_of_variables_linear:
   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
   assumes "linear g"
   shows "(\<lambda>x. \<bar>det (matrix g)\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and>
@@ -3665,14 +3665,14 @@
   qed (use h in auto)
 qed
 
-lemma%unimportant absolutely_integrable_change_of_variables_linear:
+lemma absolutely_integrable_change_of_variables_linear:
   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
   assumes "linear g"
   shows "(\<lambda>x. \<bar>det (matrix g)\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S
      \<longleftrightarrow> f absolutely_integrable_on (g ` S)"
   using assms has_absolute_integral_change_of_variables_linear by blast
 
-lemma%unimportant absolutely_integrable_on_linear_image:
+lemma absolutely_integrable_on_linear_image:
   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
   assumes "linear g"
   shows "f absolutely_integrable_on (g ` S)
@@ -3680,12 +3680,12 @@
   unfolding assms absolutely_integrable_change_of_variables_linear [OF assms, symmetric] absolutely_integrable_on_scaleR_iff
   by (auto simp: set_integrable_def)
 
-lemma%important integral_change_of_variables_linear:
+lemma integral_change_of_variables_linear:
   fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
   assumes "linear g"
       and "f absolutely_integrable_on (g ` S) \<or> (f \<circ> g) absolutely_integrable_on S"
     shows "integral (g ` S) f = \<bar>det (matrix g)\<bar> *\<^sub>R integral S (f \<circ> g)"
-proof%unimportant -
+proof -
   have "((\<lambda>x. \<bar>det (matrix g)\<bar> *\<^sub>R f (g x)) absolutely_integrable_on S) \<or> (f absolutely_integrable_on g ` S)"
     using absolutely_integrable_on_linear_image assms by blast
   moreover
@@ -3699,18 +3699,18 @@
 
 subsection\<open>Change of variable for measure\<close>
 
-lemma%unimportant has_measure_differentiable_image:
+lemma has_measure_differentiable_image:
   fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
   assumes "S \<in> sets lebesgue"
       and "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
       and "inj_on f S"
   shows "f ` S \<in> lmeasurable \<and> measure lebesgue (f ` S) = m
      \<longleftrightarrow> ((\<lambda>x. \<bar>det (matrix (f' x))\<bar>) has_integral m) S"
-  using%unimportant has_absolute_integral_change_of_variables [OF assms, of "\<lambda>x. (1::real^1)" "vec m"]
-  unfolding%unimportant absolutely_integrable_on_1_iff integral_on_1_eq integrable_on_1_iff absolutely_integrable_on_def
-  by%unimportant (auto simp: has_integral_iff lmeasurable_iff_integrable_on lmeasure_integral)
+  using has_absolute_integral_change_of_variables [OF assms, of "\<lambda>x. (1::real^1)" "vec m"]
+  unfolding absolutely_integrable_on_1_iff integral_on_1_eq integrable_on_1_iff absolutely_integrable_on_def
+  by (auto simp: has_integral_iff lmeasurable_iff_integrable_on lmeasure_integral)
 
-lemma%unimportant measurable_differentiable_image_eq:
+lemma measurable_differentiable_image_eq:
   fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
   assumes "S \<in> sets lebesgue"
       and "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
@@ -3719,23 +3719,23 @@
   using has_measure_differentiable_image [OF assms]
   by blast
 
-lemma%important measurable_differentiable_image_alt:
+lemma measurable_differentiable_image_alt:
   fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
   assumes "S \<in> sets lebesgue"
     and "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
     and "inj_on f S"
   shows "f ` S \<in> lmeasurable \<longleftrightarrow> (\<lambda>x. \<bar>det (matrix (f' x))\<bar>) absolutely_integrable_on S"
-  using%unimportant measurable_differentiable_image_eq [OF assms]
-  by%unimportant (simp only: absolutely_integrable_on_iff_nonneg)
+  using measurable_differentiable_image_eq [OF assms]
+  by (simp only: absolutely_integrable_on_iff_nonneg)
 
-lemma%important measure_differentiable_image_eq:
+lemma measure_differentiable_image_eq:
   fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
   assumes S: "S \<in> sets lebesgue"
     and der_f: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
     and inj: "inj_on f S"
     and intS: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on S"
   shows "measure lebesgue (f ` S) = integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
-  using%unimportant measurable_differentiable_image_eq [OF S der_f inj]
-        assms has_measure_differentiable_image by%unimportant blast
+  using measurable_differentiable_image_eq [OF S der_f inj]
+        assms has_measure_differentiable_image by blast
 
 end