src/HOL/Analysis/Change_Of_Vars.thy
author wenzelm
Wed, 23 Jul 2025 14:53:21 +0200
changeset 82898 89da4dcd1fa8
parent 82518 da14e77a48b2
permissions -rw-r--r--
clarified colors, following d6a14ed060fb;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
68017
e99f9b3962bf three new theorems
paulson <lp15@cam.ac.uk>
parents: 68001
diff changeset
     1
(*  Title:      HOL/Analysis/Change_Of_Vars.thy
e99f9b3962bf three new theorems
paulson <lp15@cam.ac.uk>
parents: 68001
diff changeset
     2
    Authors:    LC Paulson, based on material from HOL Light
e99f9b3962bf three new theorems
paulson <lp15@cam.ac.uk>
parents: 68001
diff changeset
     3
*)
e99f9b3962bf three new theorems
paulson <lp15@cam.ac.uk>
parents: 68001
diff changeset
     4
e99f9b3962bf three new theorems
paulson <lp15@cam.ac.uk>
parents: 68001
diff changeset
     5
section\<open>Change of Variables Theorems\<close>
e99f9b3962bf three new theorems
paulson <lp15@cam.ac.uk>
parents: 68001
diff changeset
     6
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     7
theory Change_Of_Vars
67999
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67998
diff changeset
     8
  imports Vitali_Covering_Theorem Determinants
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     9
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    10
begin
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    11
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69661
diff changeset
    12
subsection \<open>Measurable Shear and Stretch\<close>
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    13
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
    14
proposition
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    15
  fixes a :: "real^'n"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    16
  assumes "m \<noteq> n" and ab_ne: "cbox a b \<noteq> {}" and an: "0 \<le> a$n"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    17
  shows measurable_shear_interval: "(\<lambda>x. \<chi> i. if i = m then x$m + x$n else x$i) ` (cbox a b) \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    18
       (is  "?f ` _ \<in> _")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    19
   and measure_shear_interval: "measure lebesgue ((\<lambda>x. \<chi> i. if i = m then x$m + x$n else x$i) ` cbox a b)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    20
               = measure lebesgue (cbox a b)" (is "?Q")
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
    21
proof -
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    22
  have lin: "linear ?f"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 68001
diff changeset
    23
    by (rule linearI) (auto simp: plus_vec_def scaleR_vec_def algebra_simps)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    24
  show fab: "?f ` cbox a b \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    25
    by (simp add: lin measurable_linear_image_interval)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    26
  let ?c = "\<chi> i. if i = m then b$m + b$n else b$i"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    27
  let ?mn = "axis m 1 - axis n (1::real)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    28
  have eq1: "measure lebesgue (cbox a ?c)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    29
            = measure lebesgue (?f ` cbox a b)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    30
            + measure lebesgue (cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a$m})
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    31
            + measure lebesgue (cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m})"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    32
  proof (rule measure_Un3_negligible)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    33
    show "cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a$m} \<in> lmeasurable" "cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m} \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    34
      by (auto simp: convex_Int convex_halfspace_le convex_halfspace_ge bounded_Int measurable_convex)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    35
    have "negligible {x. ?mn \<bullet> x = a$m}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    36
      by (metis \<open>m \<noteq> n\<close> axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    37
    moreover have "?f ` cbox a b \<inter> (cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m}) \<subseteq> {x. ?mn \<bullet> x = a$m}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    38
      using \<open>m \<noteq> n\<close> antisym_conv by (fastforce simp: algebra_simps mem_box_cart inner_axis')
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    39
    ultimately show "negligible ((?f ` cbox a b) \<inter> (cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m}))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    40
      by (rule negligible_subset)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    41
    have "negligible {x. ?mn \<bullet> x = b$m}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    42
      by (metis \<open>m \<noteq> n\<close> axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    43
    moreover have "(?f ` cbox a b) \<inter> (cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m}) \<subseteq> {x. ?mn \<bullet> x = b$m}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    44
      using \<open>m \<noteq> n\<close> antisym_conv by (fastforce simp: algebra_simps mem_box_cart inner_axis')
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    45
    ultimately show "negligible (?f ` cbox a b \<inter> (cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m}))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    46
      by (rule negligible_subset)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    47
    have "negligible {x. ?mn \<bullet> x = b$m}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    48
      by (metis \<open>m \<noteq> n\<close> axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    49
    moreover have "(cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m} \<inter> (cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m})) \<subseteq> {x. ?mn \<bullet> x = b$m}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    50
      using \<open>m \<noteq> n\<close> ab_ne
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
    51
      apply (clarsimp simp: algebra_simps mem_box_cart inner_axis')
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
    52
      by (smt (verit, ccfv_SIG) interval_ne_empty_cart(1))
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    53
    ultimately show "negligible (cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m} \<inter> (cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m}))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    54
      by (rule negligible_subset)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    55
    show "?f ` cbox a b \<union> cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m} \<union> cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m} = cbox a ?c" (is "?lhs = _")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    56
    proof
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    57
      show "?lhs \<subseteq> cbox a ?c"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    58
        by (auto simp: mem_box_cart add_mono) (meson add_increasing2 an order_trans)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    59
      show "cbox a ?c \<subseteq> ?lhs"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
    60
        apply (clarsimp simp: algebra_simps image_iff inner_axis' lambda_add_Galois [OF \<open>m \<noteq> n\<close>])
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
    61
        by (smt (verit, del_insts) mem_box_cart(2) vec_lambda_beta)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    62
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    63
  qed (fact fab)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    64
  let ?d = "\<chi> i. if i = m then a $ m - b $ m else 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    65
  have eq2: "measure lebesgue (cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m}) + measure lebesgue (cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m})
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    66
           = measure lebesgue (cbox a (\<chi> i. if i = m then a $ m + b $ n else b $ i))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    67
  proof (rule measure_translate_add[of "cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a$m}" "cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    68
     "(\<chi> i. if i = m then a$m - b$m else 0)" "cbox a (\<chi> i. if i = m then a$m + b$n else b$i)"])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    69
    show "(cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a$m}) \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    70
      "cbox a ?c \<inter> {x. ?mn \<bullet> x \<ge> b$m} \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    71
      by (auto simp: convex_Int convex_halfspace_le convex_halfspace_ge bounded_Int measurable_convex)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    72
    have "\<And>x. \<lbrakk>x $ n + a $ m \<le> x $ m\<rbrakk>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    73
         \<Longrightarrow> x \<in> (+) (\<chi> i. if i = m then a $ m - b $ m else 0) ` {x. x $ n + b $ m \<le> x $ m}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    74
      using \<open>m \<noteq> n\<close>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    75
      by (rule_tac x="x - (\<chi> i. if i = m then a$m - b$m else 0)" in image_eqI)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    76
         (simp_all add: mem_box_cart)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    77
    then have imeq: "(+) ?d ` {x. b $ m \<le> ?mn \<bullet> x} = {x. a $ m \<le> ?mn \<bullet> x}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    78
      using \<open>m \<noteq> n\<close> by (auto simp: mem_box_cart inner_axis' algebra_simps)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    79
    have "\<And>x. \<lbrakk>0 \<le> a $ n; x $ n + a $ m \<le> x $ m;
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    80
                \<forall>i. i \<noteq> m \<longrightarrow> a $ i \<le> x $ i \<and> x $ i \<le> b $ i\<rbrakk>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    81
         \<Longrightarrow> a $ m \<le> x $ m"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    82
      using \<open>m \<noteq> n\<close>  by force
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    83
    then have "(+) ?d ` (cbox a ?c \<inter> {x. b $ m \<le> ?mn \<bullet> x})
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    84
            = cbox a (\<chi> i. if i = m then a $ m + b $ n else b $ i) \<inter> {x. a $ m \<le> ?mn \<bullet> x}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    85
      using an ab_ne
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    86
      apply (simp add: cbox_translation [symmetric] translation_Int interval_ne_empty_cart imeq)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    87
      apply (auto simp: mem_box_cart inner_axis' algebra_simps if_distrib all_if_distrib)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    88
      by (metis (full_types) add_mono mult_2_right)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    89
    then show "cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m} \<union>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    90
          (+) ?d ` (cbox a ?c \<inter> {x. b $ m \<le> ?mn \<bullet> x}) =
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    91
          cbox a (\<chi> i. if i = m then a $ m + b $ n else b $ i)"  (is "?lhs = ?rhs")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    92
      using an \<open>m \<noteq> n\<close>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    93
      apply (auto simp: mem_box_cart inner_axis' algebra_simps if_distrib all_if_distrib, force)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    94
        apply (drule_tac x=n in spec)+
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    95
      by (meson ab_ne add_mono_thms_linordered_semiring(3) dual_order.trans interval_ne_empty_cart(1))
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    96
    have "negligible{x. ?mn \<bullet> x = a$m}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    97
      by (metis \<open>m \<noteq> n\<close> axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    98
    moreover have "(cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m} \<inter>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    99
                                 (+) ?d ` (cbox a ?c \<inter> {x. b $ m \<le> ?mn \<bullet> x})) \<subseteq> {x. ?mn \<bullet> x = a$m}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   100
      using \<open>m \<noteq> n\<close> antisym_conv by (fastforce simp: algebra_simps mem_box_cart inner_axis')
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   101
    ultimately show "negligible (cbox a ?c \<inter> {x. ?mn \<bullet> x \<le> a $ m} \<inter>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   102
                                 (+) ?d ` (cbox a ?c \<inter> {x. b $ m \<le> ?mn \<bullet> x}))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   103
      by (rule negligible_subset)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   104
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   105
  have ac_ne: "cbox a ?c \<noteq> {}"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   106
    by (smt (verit, del_insts) ab_ne an interval_ne_empty_cart(1) vec_lambda_beta)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   107
  have ax_ne: "cbox a (\<chi> i. if i = m then a $ m + b $ n else b $ i) \<noteq> {}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   108
    using ab_ne an
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   109
    by (smt (verit, ccfv_threshold) interval_ne_empty_cart(1) vec_lambda_beta)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   110
  have eq3: "measure lebesgue (cbox a ?c) = measure lebesgue (cbox a (\<chi> i. if i = m then a$m + b$n else b$i)) + measure lebesgue (cbox a b)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   111
    by (simp add: content_cbox_if_cart ab_ne ac_ne ax_ne algebra_simps prod.delta_remove
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   112
             if_distrib [of "\<lambda>u. u - z" for z] prod.remove)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   113
  show ?Q
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   114
    using eq1 eq2 eq3 by (simp add: algebra_simps)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   115
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   116
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   117
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   118
proposition
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   119
  fixes S :: "(real^'n) set"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   120
  assumes "S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   121
  shows measurable_stretch: "((\<lambda>x. \<chi> k. m k * x$k) ` S) \<in> lmeasurable" (is  "?f ` S \<in> _")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   122
    and measure_stretch: "measure lebesgue ((\<lambda>x. \<chi> k. m k * x$k) ` S) = \<bar>prod m UNIV\<bar> * measure lebesgue S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   123
    (is "?MEQ")
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   124
proof -
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   125
  have "(?f ` S) \<in> lmeasurable \<and> ?MEQ"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   126
  proof (cases "\<forall>k. m k \<noteq> 0")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   127
    case True
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   128
    have m0: "0 < \<bar>prod m UNIV\<bar>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   129
      using True by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   130
    have "(indicat_real (?f ` S) has_integral \<bar>prod m UNIV\<bar> * measure lebesgue S) UNIV"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   131
    proof (clarsimp simp add: has_integral_alt [where i=UNIV])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   132
      fix e :: "real"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   133
      assume "e > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   134
      have "(indicat_real S has_integral (measure lebesgue S)) UNIV"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   135
        using assms lmeasurable_iff_has_integral by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   136
      then obtain B where "B>0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   137
        and B: "\<And>a b. ball 0 B \<subseteq> cbox a b \<Longrightarrow>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   138
                        \<exists>z. (indicat_real S has_integral z) (cbox a b) \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   139
                            \<bar>z - measure lebesgue S\<bar> < e / \<bar>prod m UNIV\<bar>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   140
        by (simp add: has_integral_alt [where i=UNIV]) (metis (full_types) divide_pos_pos m0  m0 \<open>e > 0\<close>)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   141
      show "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   142
                  (\<exists>z. (indicat_real (?f ` S) has_integral z) (cbox a b) \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   143
                       \<bar>z - \<bar>prod m UNIV\<bar> * measure lebesgue S\<bar> < e)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   144
      proof (intro exI conjI allI)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   145
        let ?C = "Max (range (\<lambda>k. \<bar>m k\<bar>)) * B"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   146
        show "?C > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   147
          using True \<open>B > 0\<close> by (simp add: Max_gr_iff)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   148
        show "ball 0 ?C \<subseteq> cbox u v \<longrightarrow>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   149
                  (\<exists>z. (indicat_real (?f ` S) has_integral z) (cbox u v) \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   150
                       \<bar>z - \<bar>prod m UNIV\<bar> * measure lebesgue S\<bar> < e)" for u v
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   151
        proof
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   152
          assume uv: "ball 0 ?C \<subseteq> cbox u v"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   153
          with \<open>?C > 0\<close> have cbox_ne: "cbox u v \<noteq> {}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   154
            using centre_in_ball by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   155
          let ?\<alpha> = "\<lambda>k. u$k / m k"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   156
          let ?\<beta> = "\<lambda>k. v$k / m k"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   157
          have invm0: "\<And>k. inverse (m k) \<noteq> 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   158
            using True by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   159
          have "ball 0 B \<subseteq> (\<lambda>x. \<chi> k. x $ k / m k) ` ball 0 ?C"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   160
          proof clarsimp
67999
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67998
diff changeset
   161
            fix x :: "real^'n"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   162
            assume x: "norm x < B"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   163
            have [simp]: "\<bar>Max (range (\<lambda>k. \<bar>m k\<bar>))\<bar> = Max (range (\<lambda>k. \<bar>m k\<bar>))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   164
              by (meson Max_ge abs_ge_zero abs_of_nonneg finite finite_imageI order_trans rangeI)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   165
            have "norm (\<chi> k. m k * x $ k) \<le> norm (Max (range (\<lambda>k. \<bar>m k\<bar>)) *\<^sub>R x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   166
              by (rule norm_le_componentwise_cart) (auto simp: abs_mult intro: mult_right_mono)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   167
            also have "\<dots> < ?C"
72569
d56e4eeae967 mult_le_cancel_iff1, mult_le_cancel_iff2, mult_less_iff1 generalised from the real_ versions
paulson <lp15@cam.ac.uk>
parents: 72445
diff changeset
   168
              using x \<open>0 < (MAX k. \<bar>m k\<bar>) * B\<close> \<open>0 < B\<close> zero_less_mult_pos2 by fastforce
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   169
            finally have "norm (\<chi> k. m k * x $ k) < ?C" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   170
            then show "x \<in> (\<lambda>x. \<chi> k. x $ k / m k) ` ball 0 ?C"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   171
              using stretch_Galois [of "inverse \<circ> m"] True by (auto simp: image_iff field_simps)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   172
          qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   173
          then have Bsub: "ball 0 B \<subseteq> cbox (\<chi> k. min (?\<alpha> k) (?\<beta> k)) (\<chi> k. max (?\<alpha> k) (?\<beta> k))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   174
            using cbox_ne uv image_stretch_interval_cart [of "inverse \<circ> m" u v, symmetric]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   175
            by (force simp: field_simps)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   176
          obtain z where zint: "(indicat_real S has_integral z) (cbox (\<chi> k. min (?\<alpha> k) (?\<beta> k)) (\<chi> k. max (?\<alpha> k) (?\<beta> k)))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   177
                   and zless: "\<bar>z - measure lebesgue S\<bar> < e / \<bar>prod m UNIV\<bar>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   178
            using B [OF Bsub] by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   179
          have ind: "indicat_real (?f ` S) = (\<lambda>x. indicator S (\<chi> k. x$k / m k))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   180
            using True stretch_Galois [of m] by (force simp: indicator_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   181
          show "\<exists>z. (indicat_real (?f ` S) has_integral z) (cbox u v) \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   182
                       \<bar>z - \<bar>prod m UNIV\<bar> * measure lebesgue S\<bar> < e"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   183
          proof (simp add: ind, intro conjI exI)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   184
            have "((\<lambda>x. indicat_real S (\<chi> k. x $ k/ m k)) has_integral z *\<^sub>R \<bar>prod m UNIV\<bar>)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   185
                ((\<lambda>x. \<chi> k. x $ k * m k) ` cbox (\<chi> k. min (?\<alpha> k) (?\<beta> k)) (\<chi> k. max (?\<alpha> k) (?\<beta> k)))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   186
              using True has_integral_stretch_cart [OF zint, of "inverse \<circ> m"]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   187
              by (simp add: field_simps prod_dividef)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   188
            moreover have "((\<lambda>x. \<chi> k. x $ k * m k) ` cbox (\<chi> k. min (?\<alpha> k) (?\<beta> k)) (\<chi> k. max (?\<alpha> k) (?\<beta> k))) = cbox u v"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   189
              using True image_stretch_interval_cart [of "inverse \<circ> m" u v, symmetric]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   190
                image_stretch_interval_cart [of "\<lambda>k. 1" u v, symmetric] \<open>cbox u v \<noteq> {}\<close>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   191
              by (simp add: field_simps image_comp o_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   192
            ultimately show "((\<lambda>x. indicat_real S (\<chi> k. x $ k/ m k)) has_integral z *\<^sub>R \<bar>prod m UNIV\<bar>) (cbox u v)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   193
              by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   194
            have "\<bar>z *\<^sub>R \<bar>prod m UNIV\<bar> - \<bar>prod m UNIV\<bar> * measure lebesgue S\<bar>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   195
                 = \<bar>prod m UNIV\<bar> * \<bar>z - measure lebesgue S\<bar>"
73932
fd21b4a93043 added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents: 73648
diff changeset
   196
              by (metis (no_types, opaque_lifting) abs_abs abs_scaleR mult.commute real_scaleR_def right_diff_distrib')
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   197
            also have "\<dots> < e"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   198
              using zless True by (simp add: field_simps)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   199
            finally show "\<bar>z *\<^sub>R \<bar>prod m UNIV\<bar> - \<bar>prod m UNIV\<bar> * measure lebesgue S\<bar> < e" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   200
          qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   201
        qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   202
      qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   203
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   204
    then show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   205
      by (auto simp: has_integral_integrable integral_unique lmeasure_integral_UNIV measurable_integrable)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   206
  next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   207
    case False
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   208
    then obtain k where "m k = 0" and prm: "prod m UNIV = 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   209
      by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   210
    have nfS: "negligible (?f ` S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   211
      by (rule negligible_subset [OF negligible_standard_hyperplane_cart]) (use \<open>m k = 0\<close> in auto)
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   212
    then show ?thesis
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   213
      by (simp add: negligible_iff_measure prm)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   214
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   215
  then show "(?f ` S) \<in> lmeasurable" ?MEQ
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   216
    by metis+
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   217
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   218
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   219
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   220
proposition
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   221
 fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   222
  assumes "linear f" "S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   223
  shows measurable_linear_image: "(f ` S) \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   224
    and measure_linear_image: "measure lebesgue (f ` S) = \<bar>det (matrix f)\<bar> * measure lebesgue S" (is "?Q f S")
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   225
proof -
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   226
  have "\<forall>S \<in> lmeasurable. (f ` S) \<in> lmeasurable \<and> ?Q f S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   227
  proof (rule induct_linear_elementary [OF \<open>linear f\<close>]; intro ballI)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   228
    fix f g and S :: "(real,'n) vec set"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   229
    assume "linear f" and "linear g"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   230
      and f [rule_format]: "\<forall>S \<in> lmeasurable. f ` S \<in> lmeasurable \<and> ?Q f S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   231
      and g [rule_format]: "\<forall>S \<in> lmeasurable. g ` S \<in> lmeasurable \<and> ?Q g S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   232
      and S: "S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   233
    then have gS: "g ` S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   234
      by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   235
    show "(f \<circ> g) ` S \<in> lmeasurable \<and> ?Q (f \<circ> g) S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   236
      using f [OF gS] g [OF S] matrix_compose [OF \<open>linear g\<close> \<open>linear f\<close>]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   237
      by (simp add: o_def image_comp abs_mult det_mul)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   238
  next
67999
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67998
diff changeset
   239
    fix f :: "real^'n::_ \<Rightarrow> real^'n::_" and i and S :: "(real^'n::_) set"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   240
    assume "linear f" and 0: "\<And>x. f x $ i = 0" and "S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   241
    then have "\<not> inj f"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   242
      by (metis (full_types) linear_injective_imp_surjective one_neq_zero surjE vec_component)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   243
    have detf: "det (matrix f) = 0"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 68001
diff changeset
   244
      using \<open>\<not> inj f\<close> det_nz_iff_inj[OF \<open>linear f\<close>] by blast
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   245
    show "f ` S \<in> lmeasurable \<and> ?Q f S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   246
    proof
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   247
      show "f ` S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   248
        using lmeasurable_iff_indicator_has_integral \<open>linear f\<close> \<open>\<not> inj f\<close> negligible_UNIV negligible_linear_singular_image by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   249
      have "measure lebesgue (f ` S) = 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   250
        by (meson \<open>\<not> inj f\<close> \<open>linear f\<close> negligible_imp_measure0 negligible_linear_singular_image)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   251
      also have "\<dots> = \<bar>det (matrix f)\<bar> * measure lebesgue S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   252
        by (simp add: detf)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   253
      finally show "?Q f S" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   254
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   255
  next
67999
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67998
diff changeset
   256
    fix c and S :: "(real^'n::_) set"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   257
    assume "S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   258
    show "(\<lambda>a. \<chi> i. c i * a $ i) ` S \<in> lmeasurable \<and> ?Q (\<lambda>a. \<chi> i. c i * a $ i) S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   259
    proof
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   260
      show "(\<lambda>a. \<chi> i. c i * a $ i) ` S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   261
        by (simp add: \<open>S \<in> lmeasurable\<close> measurable_stretch)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   262
      show "?Q (\<lambda>a. \<chi> i. c i * a $ i) S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   263
        by (simp add: measure_stretch [OF \<open>S \<in> lmeasurable\<close>, of c] axis_def matrix_def det_diagonal)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   264
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   265
  next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   266
    fix m :: "'n" and n :: "'n" and S :: "(real, 'n) vec set"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   267
    assume "m \<noteq> n" and "S \<in> lmeasurable"
73648
1bd3463e30b8 more elementary swap
haftmann
parents: 73536
diff changeset
   268
    let ?h = "\<lambda>v::(real, 'n) vec. \<chi> i. v $ Transposition.transpose m n i"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   269
    have lin: "linear ?h"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 68001
diff changeset
   270
      by (rule linearI) (simp_all add: plus_vec_def scaleR_vec_def)
73648
1bd3463e30b8 more elementary swap
haftmann
parents: 73536
diff changeset
   271
    have meq: "measure lebesgue ((\<lambda>v::(real, 'n) vec. \<chi> i. v $ Transposition.transpose m n i) ` cbox a b)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   272
             = measure lebesgue (cbox a b)" for a b
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   273
    proof (cases "cbox a b = {}")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   274
      case True then show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   275
        by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   276
    next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   277
      case False
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   278
      then have him: "?h ` (cbox a b) \<noteq> {}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   279
        by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   280
      have eq: "?h ` (cbox a b) = cbox (?h a) (?h b)"
73648
1bd3463e30b8 more elementary swap
haftmann
parents: 73536
diff changeset
   281
        by (auto simp: image_iff lambda_swap_Galois mem_box_cart) (metis transpose_involutory)+
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   282
      show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   283
        using him prod.permute [OF permutes_swap_id, where S=UNIV and g="\<lambda>i. (b - a)$i", symmetric]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   284
        by (simp add: eq content_cbox_cart False)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   285
    qed
73648
1bd3463e30b8 more elementary swap
haftmann
parents: 73536
diff changeset
   286
    have "(\<chi> i j. if Transposition.transpose m n i = j then 1 else 0) = (\<chi> i j. if j = Transposition.transpose m n i then 1 else (0::real))"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   287
      by (auto intro!: Cart_lambda_cong)
73648
1bd3463e30b8 more elementary swap
haftmann
parents: 73536
diff changeset
   288
    then have "matrix ?h = transpose(\<chi> i j. mat 1 $ i $ Transposition.transpose m n j)"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   289
      by (auto simp: matrix_eq transpose_def axis_def mat_def matrix_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   290
    then have 1: "\<bar>det (matrix ?h)\<bar> = 1"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   291
      by (simp add: det_permute_columns permutes_swap_id sign_swap_id abs_mult)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   292
    show "?h ` S \<in> lmeasurable \<and> ?Q ?h S"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   293
      using measure_linear_sufficient [OF lin \<open>S \<in> lmeasurable\<close>] meq 1 by force
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   294
  next
67999
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67998
diff changeset
   295
    fix m n :: "'n" and S :: "(real, 'n) vec set"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   296
    assume "m \<noteq> n" and "S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   297
    let ?h = "\<lambda>v::(real, 'n) vec. \<chi> i. if i = m then v $ m + v $ n else v $ i"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   298
    have lin: "linear ?h"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 68001
diff changeset
   299
      by (rule linearI) (auto simp: algebra_simps plus_vec_def scaleR_vec_def vec_eq_iff)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   300
    consider "m < n" | " n < m"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   301
      using \<open>m \<noteq> n\<close> less_linear by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   302
    then have 1: "det(matrix ?h) = 1"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   303
    proof cases
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   304
      assume "m < n"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   305
      have *: "matrix ?h $ i $ j = (0::real)" if "j < i" for i j :: 'n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   306
      proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   307
        have "axis j 1 = (\<chi> n. if n = j then 1 else (0::real))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   308
          using axis_def by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   309
        then have "(\<chi> p q. if p = m then axis q 1 $ m + axis q 1 $ n else axis q 1 $ p) $ i $ j = (0::real)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   310
          using \<open>j < i\<close> axis_def \<open>m < n\<close> by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   311
        with \<open>m < n\<close> show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   312
          by (auto simp: matrix_def axis_def cong: if_cong)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   313
      qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   314
      show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   315
        using \<open>m \<noteq> n\<close> by (subst det_upperdiagonal [OF *]) (auto simp: matrix_def axis_def cong: if_cong)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   316
    next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   317
      assume "n < m"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   318
      have *: "matrix ?h $ i $ j = (0::real)" if "j > i" for i j :: 'n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   319
      proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   320
        have "axis j 1 = (\<chi> n. if n = j then 1 else (0::real))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   321
          using axis_def by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   322
        then have "(\<chi> p q. if p = m then axis q 1 $ m + axis q 1 $ n else axis q 1 $ p) $ i $ j = (0::real)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   323
          using \<open>j > i\<close> axis_def \<open>m > n\<close> by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   324
        with \<open>m > n\<close> show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   325
          by (auto simp: matrix_def axis_def cong: if_cong)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   326
      qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   327
      show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   328
        using \<open>m \<noteq> n\<close>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   329
        by (subst det_lowerdiagonal [OF *]) (auto simp: matrix_def axis_def cong: if_cong)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   330
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   331
    have meq: "measure lebesgue (?h ` (cbox a b)) = measure lebesgue (cbox a b)" for a b
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   332
    proof (cases "cbox a b = {}")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   333
      case True then show ?thesis by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   334
    next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   335
      case False
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   336
      then have ne: "(+) (\<chi> i. if i = n then - a $ n else 0) ` cbox a b \<noteq> {}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   337
        by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   338
      let ?v = "\<chi> i. if i = n then - a $ n else 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   339
      have "?h ` cbox a b
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   340
            = (+) (\<chi> i. if i = m \<or> i = n then a $ n else 0) ` ?h ` (+) ?v ` (cbox a b)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   341
        using \<open>m \<noteq> n\<close> unfolding image_comp o_def by (force simp: vec_eq_iff)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   342
      then have "measure lebesgue (?h ` (cbox a b))
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   343
               = measure lebesgue ((\<lambda>v. \<chi> i. if i = m then v $ m + v $ n else v $ i) `
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   344
                                   (+) ?v ` cbox a b)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   345
        by (rule ssubst) (rule measure_translation)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   346
      also have "\<dots> = measure lebesgue ((\<lambda>v. \<chi> i. if i = m then v $ m + v $ n else v $ i) ` cbox (?v +a) (?v + b))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   347
        by (metis (no_types, lifting) cbox_translation)
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   348
      also have "\<dots> = measure lebesgue ((+) ?v ` cbox a b)"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   349
        apply (subst measure_shear_interval)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   350
        using \<open>m \<noteq> n\<close> ne apply auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   351
        apply (simp add: cbox_translation)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   352
        by (metis cbox_borel cbox_translation measure_completion sets_lborel)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   353
      also have "\<dots> = measure lebesgue (cbox a b)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   354
        by (rule measure_translation)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   355
        finally show ?thesis .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   356
      qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   357
    show "?h ` S \<in> lmeasurable \<and> ?Q ?h S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   358
      using measure_linear_sufficient [OF lin \<open>S \<in> lmeasurable\<close>] meq 1 by force
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   359
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   360
  with assms show "(f ` S) \<in> lmeasurable" "?Q f S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   361
    by metis+
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   362
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   363
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   364
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   365
lemma
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   366
 fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   367
  assumes f: "orthogonal_transformation f" and S: "S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   368
  shows measurable_orthogonal_image: "f ` S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   369
    and measure_orthogonal_image: "measure lebesgue (f ` S) = measure lebesgue S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   370
proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   371
  have "linear f"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   372
    by (simp add: f orthogonal_transformation_linear)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   373
  then show "f ` S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   374
    by (metis S measurable_linear_image)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   375
  show "measure lebesgue (f ` S) = measure lebesgue S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   376
    by (simp add: measure_linear_image \<open>linear f\<close> S f)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   377
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   378
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   379
proposition measure_semicontinuous_with_hausdist_explicit:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   380
  assumes "bounded S" and neg: "negligible(frontier S)" and "e > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   381
  obtains d where "d > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   382
                  "\<And>T. \<lbrakk>T \<in> lmeasurable; \<And>y. y \<in> T \<Longrightarrow> \<exists>x. x \<in> S \<and> dist x y < d\<rbrakk>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   383
                        \<Longrightarrow> measure lebesgue T < measure lebesgue S + e"
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   384
proof (cases "S = {}")
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   385
  case True
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   386
  with that \<open>e > 0\<close> show ?thesis by force
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   387
next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   388
  case False
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   389
  then have frS: "frontier S \<noteq> {}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   390
    using \<open>bounded S\<close> frontier_eq_empty not_bounded_UNIV by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   391
  have "S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   392
    by (simp add: \<open>bounded S\<close> measurable_Jordan neg)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   393
  have null: "(frontier S) \<in> null_sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   394
    by (metis neg negligible_iff_null_sets)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   395
  have "frontier S \<in> lmeasurable" and mS0: "measure lebesgue (frontier S) = 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   396
    using neg negligible_imp_measurable negligible_iff_measure by blast+
70378
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   397
  with \<open>e > 0\<close> sets_lebesgue_outer_open
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   398
  obtain U where "open U"
70378
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   399
    and U: "frontier S \<subseteq> U" "U - frontier S \<in> lmeasurable" "emeasure lebesgue (U - frontier S) < e"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   400
    by (metis fmeasurableD)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   401
  with null have "U \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   402
    by (metis borel_open measurable_Diff_null_set sets_completionI_sets sets_lborel)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   403
  have "measure lebesgue (U - frontier S) = measure lebesgue U"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   404
    using mS0 by (simp add: \<open>U \<in> lmeasurable\<close> fmeasurableD measure_Diff_null_set null)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   405
  with U have mU: "measure lebesgue U < e"
70378
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   406
    by (simp add: emeasure_eq_measure2 ennreal_less_iff)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   407
  show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   408
  proof
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   409
    have "U \<noteq> UNIV"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   410
      using \<open>U \<in> lmeasurable\<close> by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   411
    then have "- U \<noteq> {}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   412
      by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   413
    with \<open>open U\<close> \<open>frontier S \<subseteq> U\<close> show "setdist (frontier S) (- U) > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   414
      by (auto simp: \<open>bounded S\<close> open_closed compact_frontier_bounded setdist_gt_0_compact_closed frS)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   415
    fix T
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   416
    assume "T \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   417
      and T: "\<And>t. t \<in> T \<Longrightarrow> \<exists>y. y \<in> S \<and> dist y t < setdist (frontier S) (- U)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   418
    then have "measure lebesgue T - measure lebesgue S \<le> measure lebesgue (T - S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   419
      by (simp add: \<open>S \<in> lmeasurable\<close> measure_diff_le_measure_setdiff)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   420
    also have "\<dots>  \<le> measure lebesgue U"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   421
    proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   422
      have "T - S \<subseteq> U"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   423
      proof clarify
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   424
        fix x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   425
        assume "x \<in> T" and "x \<notin> S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   426
        then obtain y where "y \<in> S" and y: "dist y x < setdist (frontier S) (- U)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   427
          using T by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   428
        have "closed_segment x y \<inter> frontier S \<noteq> {}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   429
          using connected_Int_frontier \<open>x \<notin> S\<close> \<open>y \<in> S\<close> by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   430
        then obtain z where z: "z \<in> closed_segment x y" "z \<in> frontier S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   431
          by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   432
        with y have "dist z x < setdist(frontier S) (- U)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   433
          by (auto simp: dist_commute dest!: dist_in_closed_segment)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   434
        with z have False if "x \<in> -U"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   435
          using setdist_le_dist [OF \<open>z \<in> frontier S\<close> that] by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   436
        then show "x \<in> U"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   437
          by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   438
      qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   439
      then show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   440
        by (simp add: \<open>S \<in> lmeasurable\<close> \<open>T \<in> lmeasurable\<close> \<open>U \<in> lmeasurable\<close> fmeasurableD measure_mono_fmeasurable sets.Diff)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   441
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   442
    finally have "measure lebesgue T - measure lebesgue S \<le> measure lebesgue U" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   443
    with mU show "measure lebesgue T < measure lebesgue S + e"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   444
      by linarith
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   445
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   446
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   447
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   448
proposition
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   449
 fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   450
  assumes S: "S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   451
  and deriv: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   452
  and int: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   453
  and bounded: "\<And>x. x \<in> S \<Longrightarrow> \<bar>det (matrix (f' x))\<bar> \<le> B"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   454
  shows measurable_bounded_differentiable_image:
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   455
       "f ` S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   456
    and measure_bounded_differentiable_image:
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   457
       "measure lebesgue (f ` S) \<le> B * measure lebesgue S" (is "?M")
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   458
proof -
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   459
  have "f ` S \<in> lmeasurable \<and> measure lebesgue (f ` S) \<le> B * measure lebesgue S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   460
  proof (cases "B < 0")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   461
    case True
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   462
    then have "S = {}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   463
      by (meson abs_ge_zero bounded empty_iff equalityI less_le_trans linorder_not_less subsetI)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   464
    then show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   465
      by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   466
  next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   467
    case False
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   468
    then have "B \<ge> 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   469
      by arith
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   470
    let ?\<mu> = "measure lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   471
    have f_diff: "f differentiable_on S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   472
      using deriv by (auto simp: differentiable_on_def differentiable_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   473
    have eps: "f ` S \<in> lmeasurable" "?\<mu> (f ` S) \<le> (B+e) * ?\<mu> S" (is "?ME")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   474
              if "e > 0" for e
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   475
    proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   476
      have eps_d: "f ` S \<in> lmeasurable"  "?\<mu> (f ` S) \<le> (B+e) * (?\<mu> S + d)" (is "?MD")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   477
                  if "d > 0" for d
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   478
      proof -
70378
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   479
        obtain T where T: "open T" "S \<subseteq> T" and TS: "(T-S) \<in> lmeasurable" and "emeasure lebesgue (T-S) < ennreal d"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   480
          using S \<open>d > 0\<close> sets_lebesgue_outer_open by blast
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   481
        then have "?\<mu> (T-S) < d"
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   482
          by (metis emeasure_eq_measure2 ennreal_leI not_less)
ebd108578ab1 more new material about analysis
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
   483
        with S T TS have "T \<in> lmeasurable" and Tless: "?\<mu> T < ?\<mu> S + d"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   484
          by (auto simp: measurable_measure_Diff dest!: fmeasurable_Diff_D)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   485
        have "\<exists>r. 0 < r \<and> r < d \<and> ball x r \<subseteq> T \<and> f ` (S \<inter> ball x r) \<in> lmeasurable \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   486
                  ?\<mu> (f ` (S \<inter> ball x r)) \<le> (B + e) * ?\<mu> (ball x r)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   487
          if "x \<in> S" "d > 0" for x d
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   488
        proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   489
          have lin: "linear (f' x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   490
            and lim0: "((\<lambda>y. (f y - (f x + f' x (y - x))) /\<^sub>R norm(y - x)) \<longlongrightarrow> 0) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   491
            using deriv \<open>x \<in> S\<close> by (auto simp: has_derivative_within bounded_linear.linear field_simps)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   492
          have bo: "bounded (f' x ` ball 0 1)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   493
            by (simp add: bounded_linear_image linear_linear lin)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   494
          have neg: "negligible (frontier (f' x ` ball 0 1))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   495
            using deriv has_derivative_linear \<open>x \<in> S\<close>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   496
            by (auto intro!: negligible_convex_frontier [OF convex_linear_image])
71192
a8ccea88b725 Flattened dependency tree of HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71189
diff changeset
   497
          let ?unit_vol = "content (ball (0 :: real ^ 'n :: {finite, wellorder}) 1)"
a8ccea88b725 Flattened dependency tree of HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71189
diff changeset
   498
          have 0: "0 < e * ?unit_vol"
a8ccea88b725 Flattened dependency tree of HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71189
diff changeset
   499
            using \<open>e > 0\<close> by (simp add: content_ball_pos)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   500
          obtain k where "k > 0" and k:
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   501
                  "\<And>U. \<lbrakk>U \<in> lmeasurable; \<And>y. y \<in> U \<Longrightarrow> \<exists>z. z \<in> f' x ` ball 0 1 \<and> dist z y < k\<rbrakk>
71192
a8ccea88b725 Flattened dependency tree of HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71189
diff changeset
   502
                        \<Longrightarrow> ?\<mu> U < ?\<mu> (f' x ` ball 0 1) + e * ?unit_vol"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   503
            using measure_semicontinuous_with_hausdist_explicit [OF bo neg 0] by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   504
          obtain l where "l > 0" and l: "ball x l \<subseteq> T"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   505
            using \<open>x \<in> S\<close> \<open>open T\<close> \<open>S \<subseteq> T\<close> openE by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   506
          obtain \<zeta> where "0 < \<zeta>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   507
            and \<zeta>: "\<And>y. \<lbrakk>y \<in> S; y \<noteq> x; dist y x < \<zeta>\<rbrakk>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   508
                        \<Longrightarrow> norm (f y - (f x + f' x (y - x))) / norm (y - x) < k"
70802
160eaf566bcb formally augmented corresponding rules for field_simps
haftmann
parents: 70724
diff changeset
   509
            using lim0 \<open>k > 0\<close> by (simp add: Lim_within) (auto simp add: field_simps)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   510
          define r where "r \<equiv> min (min l (\<zeta>/2)) (min 1 (d/2))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   511
          show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   512
          proof (intro exI conjI)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   513
            show "r > 0" "r < d"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   514
              using \<open>l > 0\<close> \<open>\<zeta> > 0\<close> \<open>d > 0\<close> by (auto simp: r_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   515
            have "r \<le> l"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   516
              by (auto simp: r_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   517
            with l show "ball x r \<subseteq> T"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   518
              by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   519
            have ex_lessK: "\<exists>x' \<in> ball 0 1. dist (f' x x') ((f y - f x) /\<^sub>R r) < k"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   520
              if "y \<in> S" and "dist x y < r" for y
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   521
            proof (cases "y = x")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   522
              case True
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   523
              with lin linear_0 \<open>k > 0\<close> that show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   524
                by (rule_tac x=0 in bexI) (auto simp: linear_0)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   525
            next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   526
              case False
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   527
              then show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   528
              proof (rule_tac x="(y - x) /\<^sub>R r" in bexI)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   529
                have "f' x ((y - x) /\<^sub>R r) = f' x (y - x) /\<^sub>R r"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 68001
diff changeset
   530
                  by (simp add: lin linear_scale)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   531
                then have "dist (f' x ((y - x) /\<^sub>R r)) ((f y - f x) /\<^sub>R r) = norm (f' x (y - x) /\<^sub>R r - (f y - f x) /\<^sub>R r)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   532
                  by (simp add: dist_norm)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   533
                also have "\<dots> = norm (f' x (y - x) - (f y - f x)) / r"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70802
diff changeset
   534
                  using \<open>r > 0\<close> by (simp add: divide_simps scale_right_diff_distrib [symmetric])
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   535
                also have "\<dots> \<le> norm (f y - (f x + f' x (y - x))) / norm (y - x)"
71633
07bec530f02e cleaned proofs
nipkow
parents: 71192
diff changeset
   536
                  using that \<open>r > 0\<close> False by (simp add: field_split_simps dist_norm norm_minus_commute mult_right_mono)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   537
                also have "\<dots> < k"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   538
                  using that \<open>0 < \<zeta>\<close> by (simp add: dist_commute r_def  \<zeta> [OF \<open>y \<in> S\<close> False])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   539
                finally show "dist (f' x ((y - x) /\<^sub>R r)) ((f y - f x) /\<^sub>R r) < k" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   540
                show "(y - x) /\<^sub>R r \<in> ball 0 1"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   541
                  using that \<open>r > 0\<close> by (simp add: dist_norm divide_simps norm_minus_commute)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   542
              qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   543
            qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   544
            let ?rfs = "(\<lambda>x. x /\<^sub>R r) ` (+) (- f x) ` f ` (S \<inter> ball x r)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   545
            have rfs_mble: "?rfs \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   546
            proof (rule bounded_set_imp_lmeasurable)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   547
              have "f differentiable_on S \<inter> ball x r"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   548
                using f_diff by (auto simp: fmeasurableD differentiable_on_subset)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   549
              with S show "?rfs \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   550
                by (auto simp: sets.Int intro!: lebesgue_sets_translation differentiable_image_in_sets_lebesgue)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   551
              let ?B = "(\<lambda>(x, y). x + y) ` (f' x ` ball 0 1 \<times> ball 0 k)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   552
              have "bounded ?B"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   553
                by (simp add: bounded_plus [OF bo])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   554
              moreover have "?rfs \<subseteq> ?B"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   555
                apply (auto simp: dist_norm image_iff dest!: ex_lessK)
73932
fd21b4a93043 added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents: 73648
diff changeset
   556
                by (metis (no_types, opaque_lifting) add.commute diff_add_cancel dist_0_norm dist_commute dist_norm mem_ball)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   557
              ultimately show "bounded (?rfs)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   558
                by (rule bounded_subset)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   559
            qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   560
            then have "(\<lambda>x. r *\<^sub>R x) ` ?rfs \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   561
              by (simp add: measurable_linear_image)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   562
            with \<open>r > 0\<close> have "(+) (- f x) ` f ` (S \<inter> ball x r) \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   563
              by (simp add: image_comp o_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   564
            then have "(+) (f x) ` (+) (- f x) ` f ` (S \<inter> ball x r) \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   565
              using  measurable_translation by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   566
            then show fsb: "f ` (S \<inter> ball x r) \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   567
              by (simp add: image_comp o_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   568
            have "?\<mu> (f ` (S \<inter> ball x r)) = ?\<mu> (?rfs) * r ^ CARD('n)"
69661
a03a63b81f44 tuned proofs
haftmann
parents: 69325
diff changeset
   569
              using \<open>r > 0\<close> fsb
a03a63b81f44 tuned proofs
haftmann
parents: 69325
diff changeset
   570
              by (simp add: measure_linear_image measure_translation_subtract measurable_translation_subtract field_simps cong: image_cong_simp)
71192
a8ccea88b725 Flattened dependency tree of HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71189
diff changeset
   571
            also have "\<dots> \<le> (\<bar>det (matrix (f' x))\<bar> * ?unit_vol + e * ?unit_vol) * r ^ CARD('n)"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   572
            proof -
71192
a8ccea88b725 Flattened dependency tree of HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71189
diff changeset
   573
              have "?\<mu> (?rfs) < ?\<mu> (f' x ` ball 0 1) + e * ?unit_vol"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   574
                using rfs_mble by (force intro: k dest!: ex_lessK)
71192
a8ccea88b725 Flattened dependency tree of HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71189
diff changeset
   575
              then have "?\<mu> (?rfs) < \<bar>det (matrix (f' x))\<bar> * ?unit_vol + e * ?unit_vol"
a8ccea88b725 Flattened dependency tree of HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71189
diff changeset
   576
                by (simp add: lin measure_linear_image [of "f' x"])
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   577
              with \<open>r > 0\<close> show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   578
                by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   579
            qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   580
            also have "\<dots> \<le> (B + e) * ?\<mu> (ball x r)"
71192
a8ccea88b725 Flattened dependency tree of HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71189
diff changeset
   581
              using bounded [OF \<open>x \<in> S\<close>] \<open>r > 0\<close>
a8ccea88b725 Flattened dependency tree of HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71189
diff changeset
   582
              by (simp add: algebra_simps content_ball_conv_unit_ball[of r] content_ball_pos)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   583
            finally show "?\<mu> (f ` (S \<inter> ball x r)) \<le> (B + e) * ?\<mu> (ball x r)" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   584
          qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   585
        qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   586
        then obtain r where
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   587
          r0d: "\<And>x d. \<lbrakk>x \<in> S; d > 0\<rbrakk> \<Longrightarrow> 0 < r x d \<and> r x d < d"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   588
          and rT: "\<And>x d. \<lbrakk>x \<in> S; d > 0\<rbrakk> \<Longrightarrow> ball x (r x d) \<subseteq> T"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   589
          and r: "\<And>x d. \<lbrakk>x \<in> S; d > 0\<rbrakk> \<Longrightarrow>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   590
                  (f ` (S \<inter> ball x (r x d))) \<in> lmeasurable \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   591
                  ?\<mu> (f ` (S \<inter> ball x (r x d))) \<le> (B + e) * ?\<mu> (ball x (r x d))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   592
          by metis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   593
        obtain C where "countable C" and Csub: "C \<subseteq> {(x,r x t) |x t. x \<in> S \<and> 0 < t}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   594
          and pwC: "pairwise (\<lambda>i j. disjnt (ball (fst i) (snd i)) (ball (fst j) (snd j))) C"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   595
          and negC: "negligible(S - (\<Union>i \<in> C. ball (fst i) (snd i)))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   596
          apply (rule Vitali_covering_theorem_balls [of S "{(x,r x t) |x t. x \<in> S \<and> 0 < t}" fst snd])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   597
           apply auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   598
          by (metis dist_eq_0_iff r0d)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   599
        let ?UB = "(\<Union>(x,s) \<in> C. ball x s)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   600
        have eq: "f ` (S \<inter> ?UB) = (\<Union>(x,s) \<in> C. f ` (S \<inter> ball x s))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   601
          by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   602
        have mle: "?\<mu> (\<Union>(x,s) \<in> K. f ` (S \<inter> ball x s)) \<le> (B + e) * (?\<mu> S + d)"  (is "?l \<le> ?r")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   603
          if "K \<subseteq> C" and "finite K" for K
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   604
        proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   605
          have gt0: "b > 0" if "(a, b) \<in> K" for a b
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   606
            using Csub that \<open>K \<subseteq> C\<close> r0d by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   607
          have inj: "inj_on (\<lambda>(x, y). ball x y) K"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   608
            by (force simp: inj_on_def ball_eq_ball_iff dest: gt0)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   609
          have disjnt: "disjoint ((\<lambda>(x, y). ball x y) ` K)"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   610
            using pwC that pairwise_image pairwise_mono by fastforce
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   611
          have "?l \<le> (\<Sum>i\<in>K. ?\<mu> (case i of (x, s) \<Rightarrow> f ` (S \<inter> ball x s)))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   612
          proof (rule measure_UNION_le [OF \<open>finite K\<close>], clarify)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   613
            fix x r
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   614
            assume "(x,r) \<in> K"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   615
            then have "x \<in> S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   616
              using Csub \<open>K \<subseteq> C\<close> by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   617
            show "f ` (S \<inter> ball x r) \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   618
              by (meson Int_lower1 S differentiable_on_subset f_diff fmeasurableD lmeasurable_ball order_refl sets.Int differentiable_image_in_sets_lebesgue)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   619
          qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   620
          also have "\<dots> \<le> (\<Sum>(x,s) \<in> K. (B + e) * ?\<mu> (ball x s))"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   621
            using Csub r \<open>K \<subseteq> C\<close>  by (intro sum_mono) auto
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   622
          also have "\<dots> = (B + e) * (\<Sum>(x,s) \<in> K. ?\<mu> (ball x s))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   623
            by (simp add: prod.case_distrib sum_distrib_left)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   624
          also have "\<dots> = (B + e) * sum ?\<mu> ((\<lambda>(x, y). ball x y) ` K)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   625
            using \<open>B \<ge> 0\<close> \<open>e > 0\<close> by (simp add: inj sum.reindex prod.case_distrib)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   626
          also have "\<dots> = (B + e) * ?\<mu> (\<Union>(x,s) \<in> K. ball x s)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   627
            using \<open>B \<ge> 0\<close> \<open>e > 0\<close> that
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   628
            by (subst measure_Union') (auto simp: disjnt measure_Union')
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   629
          also have "\<dots> \<le> (B + e) * ?\<mu> T"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   630
            using \<open>B \<ge> 0\<close> \<open>e > 0\<close> that apply simp
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   631
            using measure_mono_fmeasurable [OF _ _ \<open>T \<in> lmeasurable\<close>] Csub rT
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   632
            by (smt (verit) SUP_least measure_nonneg measure_notin_sets mem_Collect_eq old.prod.case subset_iff)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   633
          also have "\<dots> \<le> (B + e) * (?\<mu> S + d)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   634
            using \<open>B \<ge> 0\<close> \<open>e > 0\<close> Tless by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   635
          finally show ?thesis .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   636
        qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   637
        have fSUB_mble: "(f ` (S \<inter> ?UB)) \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   638
          unfolding eq using Csub r False \<open>e > 0\<close> that
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   639
          by (auto simp: intro!: fmeasurable_UN_bound [OF \<open>countable C\<close> _ mle])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   640
        have fSUB_meas: "?\<mu> (f ` (S \<inter> ?UB)) \<le> (B + e) * (?\<mu> S + d)"  (is "?MUB")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   641
          unfolding eq using Csub r False \<open>e > 0\<close> that
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   642
          by (auto simp: intro!: measure_UN_bound [OF \<open>countable C\<close> _ mle])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   643
        have neg: "negligible ((f ` (S \<inter> ?UB) - f ` S) \<union> (f ` S - f ` (S \<inter> ?UB)))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   644
        proof (rule negligible_subset [OF negligible_differentiable_image_negligible [OF order_refl negC, where f=f]])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   645
          show "f differentiable_on S - (\<Union>i\<in>C. ball (fst i) (snd i))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   646
            by (meson DiffE differentiable_on_subset subsetI f_diff)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   647
        qed force
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   648
        show "f ` S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   649
          by (rule lmeasurable_negligible_symdiff [OF fSUB_mble neg])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   650
        show ?MD
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   651
          using fSUB_meas measure_negligible_symdiff [OF fSUB_mble neg] by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   652
      qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   653
      show "f ` S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   654
        using eps_d [of 1] by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   655
      show ?ME
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   656
      proof (rule field_le_epsilon)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   657
        fix \<delta> :: real
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   658
        assume "0 < \<delta>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   659
        then show "?\<mu> (f ` S) \<le> (B + e) * ?\<mu> S + \<delta>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   660
          using eps_d [of "\<delta> / (B+e)"] \<open>e > 0\<close> \<open>B \<ge> 0\<close> by (auto simp: divide_simps mult_ac)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   661
      qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   662
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   663
    show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   664
    proof (cases "?\<mu> S = 0")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   665
      case True
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   666
      with eps have "?\<mu> (f ` S) = 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   667
        by (metis mult_zero_right not_le zero_less_measure_iff)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   668
      then show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   669
        using eps [of 1] by (simp add: True)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   670
    next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   671
      case False
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   672
      have "?\<mu> (f ` S) \<le> B * ?\<mu> S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   673
      proof (rule field_le_epsilon)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   674
        fix e :: real
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   675
        assume "e > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   676
        then show "?\<mu> (f ` S) \<le> B * ?\<mu> S + e"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   677
          using eps [of "e / ?\<mu> S"] False by (auto simp: algebra_simps zero_less_measure_iff)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   678
      qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   679
      with eps [of 1] show ?thesis by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   680
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   681
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   682
  then show "f ` S \<in> lmeasurable" ?M by blast+
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   683
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   684
69739
nipkow
parents: 69720
diff changeset
   685
lemma m_diff_image_weak:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   686
 fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   687
  assumes S: "S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   688
    and deriv: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   689
    and int: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on S"
69739
nipkow
parents: 69720
diff changeset
   690
  shows "f ` S \<in> lmeasurable \<and> measure lebesgue (f ` S) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   691
proof -
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   692
  let ?\<mu> = "measure lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   693
  have aint_S: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) absolutely_integrable_on S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   694
    using int unfolding absolutely_integrable_on_def by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   695
  define m where "m \<equiv> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   696
  have *: "f ` S \<in> lmeasurable" "?\<mu> (f ` S) \<le> m + e * ?\<mu> S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   697
    if "e > 0" for e
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   698
  proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   699
    define T where "T \<equiv> \<lambda>n. {x \<in> S. n * e \<le> \<bar>det (matrix (f' x))\<bar> \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   700
                                     \<bar>det (matrix (f' x))\<bar> < (Suc n) * e}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   701
    have meas_t: "T n \<in> lmeasurable" for n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   702
    proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   703
      have *: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) \<in> borel_measurable (lebesgue_on S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   704
        using aint_S by (simp add: S borel_measurable_restrict_space_iff fmeasurableD set_integrable_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   705
      have [intro]: "x \<in> sets (lebesgue_on S) \<Longrightarrow> x \<in> sets lebesgue" for x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   706
        using S sets_restrict_space_subset by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   707
      have "{x \<in> S. real n * e \<le> \<bar>det (matrix (f' x))\<bar>} \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   708
        using * by (auto simp: borel_measurable_iff_halfspace_ge space_restrict_space)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   709
      then have 1: "{x \<in> S. real n * e \<le> \<bar>det (matrix (f' x))\<bar>} \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   710
        using S by (simp add: fmeasurableI2)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   711
      have "{x \<in> S. \<bar>det (matrix (f' x))\<bar> < (1 + real n) * e} \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   712
        using * by (auto simp: borel_measurable_iff_halfspace_less space_restrict_space)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   713
      then have 2: "{x \<in> S. \<bar>det (matrix (f' x))\<bar> < (1 + real n) * e} \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   714
        using S by (simp add: fmeasurableI2)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   715
      show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   716
        using fmeasurable.Int [OF 1 2] by (simp add: T_def Int_def cong: conj_cong)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   717
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   718
    have aint_T: "\<And>k. (\<lambda>x. \<bar>det (matrix (f' x))\<bar>) absolutely_integrable_on T k"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   719
      using set_integrable_subset [OF aint_S] meas_t T_def by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   720
    have Seq: "S = (\<Union>n. T n)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   721
      apply (auto simp: T_def)
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   722
      apply (rule_tac x = "nat \<lfloor>\<bar>det (matrix (f' x))\<bar> / e\<rfloor>" in exI)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   723
      by (smt (verit, del_insts) divide_nonneg_nonneg floor_eq_iff of_nat_nat pos_divide_less_eq that zero_le_floor)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   724
    have meas_ft: "f ` T n \<in> lmeasurable" for n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   725
    proof (rule measurable_bounded_differentiable_image)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   726
      show "T n \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   727
        by (simp add: meas_t)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   728
    next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   729
      fix x :: "(real,'n) vec"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   730
      assume "x \<in> T n"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   731
      show "(f has_derivative f' x) (at x within T n)"
72445
2c2de074832e tidying and removal of legacy name
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   732
        by (metis (no_types, lifting) \<open>x \<in> T n\<close> deriv has_derivative_subset mem_Collect_eq subsetI T_def)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   733
      show "\<bar>det (matrix (f' x))\<bar> \<le> (Suc n) * e"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   734
        using \<open>x \<in> T n\<close> T_def by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   735
    next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   736
      show "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on T n"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   737
        using aint_T absolutely_integrable_on_def by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   738
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   739
    have disT: "disjoint (range T)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   740
      unfolding disjoint_def
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   741
    proof clarsimp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   742
      show "T m \<inter> T n = {}" if "T m \<noteq> T n" for m n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   743
        using that
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   744
      proof (induction m n rule: linorder_less_wlog)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   745
        case (less m n)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   746
        with \<open>e > 0\<close> show ?case
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   747
          unfolding T_def
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   748
          proof (clarsimp simp add: Collect_conj_eq [symmetric])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   749
            fix x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   750
            assume "e > 0"  "m < n"  "n * e \<le> \<bar>det (matrix (f' x))\<bar>"  "\<bar>det (matrix (f' x))\<bar> < (1 + real m) * e"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   751
            then have "n < 1 + real m"
79566
f783490c6c99 A small number of new lemmas
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
   752
              by (metis (no_types, opaque_lifting) less_le_trans mult.commute not_le mult_le_cancel_left_pos)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   753
            then show "False"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   754
              using less.hyps by linarith
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   755
          qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   756
      qed auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   757
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   758
    have injT: "inj_on T ({n. T n \<noteq> {}})"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   759
      unfolding inj_on_def
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   760
    proof clarsimp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   761
      show "m = n" if "T m = T n" "T n \<noteq> {}" for m n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   762
        using that
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   763
      proof (induction m n rule: linorder_less_wlog)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   764
        case (less m n)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   765
        have False if "T n \<subseteq> T m" "x \<in> T n" for x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   766
          using \<open>e > 0\<close> \<open>m < n\<close> that
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   767
          apply (auto simp: T_def mult.commute intro: less_le_trans dest!: subsetD)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   768
          by (smt (verit, best) mult_less_cancel_left_disj nat_less_real_le)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   769
        then show ?case
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   770
          using less.prems by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   771
      qed auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   772
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   773
    have sum_eq_Tim: "(\<Sum>k\<le>n. f (T k)) = sum f (T ` {..n})" if "f {} = 0" for f :: "_ \<Rightarrow> real" and n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   774
    proof (subst sum.reindex_nontrivial)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   775
      fix i j  assume "i \<in> {..n}" "j \<in> {..n}" "i \<noteq> j" "T i = T j"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   776
      with that  injT [unfolded inj_on_def] show "f (T i) = 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   777
        by simp metis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   778
    qed (use atMost_atLeast0 in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   779
    let ?B = "m + e * ?\<mu> S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   780
    have "(\<Sum>k\<le>n. ?\<mu> (f ` T k)) \<le> ?B" for n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   781
    proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   782
      have "(\<Sum>k\<le>n. ?\<mu> (f ` T k)) \<le> (\<Sum>k\<le>n. ((k+1) * e) * ?\<mu>(T k))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   783
      proof (rule sum_mono [OF measure_bounded_differentiable_image])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   784
        show "(f has_derivative f' x) (at x within T k)" if "x \<in> T k" for k x
72445
2c2de074832e tidying and removal of legacy name
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   785
          using that unfolding T_def by (blast intro: deriv has_derivative_subset)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   786
        show "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on T k" for k
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   787
          using absolutely_integrable_on_def aint_T by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   788
        show "\<bar>det (matrix (f' x))\<bar> \<le> real (k + 1) * e" if "x \<in> T k" for k x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   789
          using T_def that by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   790
      qed (use meas_t in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   791
      also have "\<dots> \<le> (\<Sum>k\<le>n. (k * e) * ?\<mu>(T k)) + (\<Sum>k\<le>n. e * ?\<mu>(T k))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   792
        by (simp add: algebra_simps sum.distrib)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   793
      also have "\<dots> \<le> ?B"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   794
      proof (rule add_mono)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   795
        have "(\<Sum>k\<le>n. real k * e * ?\<mu> (T k)) = (\<Sum>k\<le>n. integral (T k) (\<lambda>x. k * e))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   796
          by (simp add: lmeasure_integral [OF meas_t]
68403
223172b97d0b reorient -> split; documented split
nipkow
parents: 68074
diff changeset
   797
                   flip: integral_mult_right integral_mult_left)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   798
        also have "\<dots> \<le> (\<Sum>k\<le>n. integral (T k) (\<lambda>x.  (abs (det (matrix (f' x))))))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   799
        proof (rule sum_mono)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   800
          fix k
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   801
          assume "k \<in> {..n}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   802
          show "integral (T k) (\<lambda>x. k * e) \<le> integral (T k) (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   803
          proof (rule integral_le [OF integrable_on_const [OF meas_t]])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   804
            show "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on T k"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   805
              using absolutely_integrable_on_def aint_T by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   806
          next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   807
            fix x assume "x \<in> T k"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   808
            show "k * e \<le> \<bar>det (matrix (f' x))\<bar>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   809
              using \<open>x \<in> T k\<close> T_def by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   810
          qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   811
        qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   812
        also have "\<dots> = sum (\<lambda>T. integral T (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)) (T ` {..n})"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   813
          by (auto intro: sum_eq_Tim)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   814
        also have "\<dots> = integral (\<Union>k\<le>n. T k) (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   815
        proof (rule integral_unique [OF has_integral_Union, symmetric])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   816
          fix S  assume "S \<in> T ` {..n}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   817
          then show "((\<lambda>x. \<bar>det (matrix (f' x))\<bar>) has_integral integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)) S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   818
          using absolutely_integrable_on_def aint_T by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   819
        next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   820
          show "pairwise (\<lambda>S S'. negligible (S \<inter> S')) (T ` {..n})"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   821
            using disT unfolding disjnt_iff by (auto simp: pairwise_def intro!: empty_imp_negligible)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   822
        qed auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   823
        also have "\<dots> \<le> m"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   824
          unfolding m_def
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   825
        proof (rule integral_subset_le)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   826
          have "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) absolutely_integrable_on (\<Union>k\<le>n. T k)"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   827
          proof (rule set_integrable_subset [OF aint_S])
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   828
            show "\<Union> (T ` {..n}) \<in> sets lebesgue"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   829
              by (intro measurable meas_t fmeasurableD)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   830
          qed (force simp: Seq)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   831
          then show "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on (\<Union>k\<le>n. T k)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   832
            using absolutely_integrable_on_def by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   833
        qed (use Seq int in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   834
        finally show "(\<Sum>k\<le>n. real k * e * ?\<mu> (T k)) \<le> m" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   835
      next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   836
        have "(\<Sum>k\<le>n. ?\<mu> (T k)) = sum ?\<mu> (T ` {..n})"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   837
          by (auto intro: sum_eq_Tim)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   838
        also have "\<dots> = ?\<mu> (\<Union>k\<le>n. T k)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   839
          using S disT by (auto simp: pairwise_def meas_t intro: measure_Union' [symmetric])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   840
        also have "\<dots> \<le> ?\<mu> S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   841
          using S by (auto simp: Seq intro: meas_t fmeasurableD measure_mono_fmeasurable)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   842
        finally have "(\<Sum>k\<le>n. ?\<mu> (T k)) \<le> ?\<mu> S" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   843
        then show "(\<Sum>k\<le>n. e * ?\<mu> (T k)) \<le> e * ?\<mu> S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   844
          by (metis less_eq_real_def ordered_comm_semiring_class.comm_mult_left_mono sum_distrib_left that)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   845
      qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   846
      finally show "(\<Sum>k\<le>n. ?\<mu> (f ` T k)) \<le> ?B" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   847
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   848
    moreover have "measure lebesgue (\<Union>k\<le>n. f ` T k) \<le> (\<Sum>k\<le>n. ?\<mu> (f ` T k))" for n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   849
      by (simp add: fmeasurableD meas_ft measure_UNION_le)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   850
    ultimately have B_ge_m: "?\<mu> (\<Union>k\<le>n. (f ` T k)) \<le> ?B" for n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   851
      by (meson order_trans)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   852
    have "(\<Union>n. f ` T n) \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   853
      by (rule fmeasurable_countable_Union [OF meas_ft B_ge_m])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   854
    moreover have "?\<mu> (\<Union>n. f ` T n) \<le> m + e * ?\<mu> S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   855
      by (rule measure_countable_Union_le [OF meas_ft B_ge_m])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   856
    ultimately show "f ` S \<in> lmeasurable" "?\<mu> (f ` S) \<le> m + e * ?\<mu> S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   857
      by (auto simp: Seq image_Union)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   858
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   859
  show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   860
  proof
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   861
    show "f ` S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   862
      using * linordered_field_no_ub by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   863
    let ?x = "m - ?\<mu> (f ` S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   864
    have False if "?\<mu> (f ` S) > integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   865
    proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   866
      have ml: "m < ?\<mu> (f ` S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   867
        using m_def that by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   868
      then have "?\<mu> S \<noteq> 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   869
        using "*"(2) bgauge_existence_lemma by fastforce
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   870
      with ml have 0: "0 < - (m - ?\<mu> (f ` S))/2 / ?\<mu> S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   871
        using that zero_less_measure_iff by force
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   872
      then show ?thesis
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70802
diff changeset
   873
        using * [OF 0] that by (auto simp: field_split_simps m_def split: if_split_asm)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   874
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   875
    then show "?\<mu> (f ` S) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   876
      by fastforce
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   877
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   878
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   879
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   880
69739
nipkow
parents: 69720
diff changeset
   881
theorem
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   882
 fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   883
  assumes S: "S \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   884
    and deriv: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   885
    and int: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   886
  shows measurable_differentiable_image: "f ` S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   887
    and measure_differentiable_image:
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   888
       "measure lebesgue (f ` S) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)" (is "?M")
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   889
proof -
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   890
  let ?I = "\<lambda>n::nat. cbox (vec (-n)) (vec n) \<inter> S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   891
  let ?\<mu> = "measure lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   892
  have "x \<in> cbox (vec (- real (nat \<lceil>norm x\<rceil>))) (vec (real (nat \<lceil>norm x\<rceil>)))" for x :: "real^'n::_"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   893
    apply (simp add: mem_box_cart)
79945
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 79566
diff changeset
   894
    by (smt (verit, best) component_le_norm_cart le_of_int_ceiling)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   895
  then have Seq: "S = (\<Union>n. ?I n)"
79945
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 79566
diff changeset
   896
    by blast
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   897
  have fIn: "f ` ?I n \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   898
       and mfIn: "?\<mu> (f ` ?I n) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)" (is ?MN) for n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   899
  proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   900
    have In: "?I n \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   901
      by (simp add: S bounded_Int bounded_set_imp_lmeasurable sets.Int)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   902
    moreover have "\<And>x. x \<in> ?I n \<Longrightarrow> (f has_derivative f' x) (at x within ?I n)"
72445
2c2de074832e tidying and removal of legacy name
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   903
      by (meson Int_iff deriv has_derivative_subset subsetI)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   904
    moreover have int_In: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on ?I n"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   905
      by (metis (mono_tags) Int_commute int integrable_altD(1) integrable_restrict_Int)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   906
    ultimately have "f ` ?I n \<in> lmeasurable" "?\<mu> (f ` ?I n) \<le> integral (?I n) (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   907
      using m_diff_image_weak by metis+
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   908
    moreover have "integral (?I n) (\<lambda>x. \<bar>det (matrix (f' x))\<bar>) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   909
      by (simp add: int_In int integral_subset_le)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   910
    ultimately show "f ` ?I n \<in> lmeasurable" ?MN
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   911
      by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   912
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   913
  have "?I k \<subseteq> ?I n" if "k \<le> n" for k n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   914
    by (rule Int_mono) (use that in \<open>auto simp: subset_interval_imp_cart\<close>)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   915
  then have "(\<Union>k\<le>n. f ` ?I k) = f ` ?I n" for n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   916
    by (fastforce simp add:)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   917
  with mfIn have "?\<mu> (\<Union>k\<le>n. f ` ?I k) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)" for n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   918
    by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   919
  then have "(\<Union>n. f ` ?I n) \<in> lmeasurable" "?\<mu> (\<Union>n. f ` ?I n) \<le> integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   920
    by (rule fmeasurable_countable_Union [OF fIn] measure_countable_Union_le [OF fIn])+
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   921
  then show "f ` S \<in> lmeasurable" ?M
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   922
    by (metis Seq image_UN)+
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   923
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   924
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   925
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   926
lemma borel_measurable_simple_function_limit_increasing:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   927
  fixes f :: "'a::euclidean_space \<Rightarrow> real"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   928
  shows "(f \<in> borel_measurable lebesgue \<and> (\<forall>x. 0 \<le> f x)) \<longleftrightarrow>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   929
         (\<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>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   930
              (\<forall>n. g n \<in> borel_measurable lebesgue) \<and> (\<forall>n. finite(range (g n))) \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   931
              (\<forall>x. (\<lambda>n. g n x) \<longlonglongrightarrow> f x))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   932
         (is "?lhs = ?rhs")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   933
proof
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   934
  assume f: ?lhs
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   935
  have leb_f: "{x. a \<le> f x \<and> f x < b} \<in> sets lebesgue" for a b
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   936
  proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   937
    have "{x. a \<le> f x \<and> f x < b} = {x. f x < b} - {x. f x < a}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   938
      by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   939
    also have "\<dots> \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   940
      using borel_measurable_vimage_halfspace_component_lt [of f UNIV] f by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   941
    finally show ?thesis .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   942
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   943
  have "g n x \<le> f x"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   944
        if inc_g: "\<And>n x. 0 \<le> g n x \<and> g n x \<le> g (Suc n) x"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   945
           and meas_g: "\<And>n. g n \<in> borel_measurable lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   946
           and fin: "\<And>n. finite(range (g n))" and lim: "\<And>x. (\<lambda>n. g n x) \<longlonglongrightarrow> f x" for g n x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   947
  proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   948
    have "\<exists>r>0. \<forall>N. \<exists>n\<ge>N. dist (g n x) (f x) \<ge> r" if "g n x > f x"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   949
    proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   950
      have g: "g n x \<le> g (N + n) x" for N
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   951
        by (rule transitive_stepwise_le) (use inc_g in auto)
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   952
      have "\<exists>m\<ge>N. g n x - f x \<le> dist (g m x) (f x)" for N
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   953
      proof
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   954
        show "N \<le> N + n \<and> g n x - f x \<le> dist (g (N + n) x) (f x)"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   955
          using g [of N] by (auto simp: dist_norm)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   956
      qed
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   957
      with that show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   958
        using diff_gt_0_iff_gt by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   959
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   960
    with lim show ?thesis
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   961
      unfolding lim_sequentially
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   962
      by (meson less_le_not_le not_le_imp_less)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   963
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   964
  moreover
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   965
  let ?\<Omega> = "\<lambda>n k. indicator {y. k/2^n \<le> f y \<and> f y < (k+1)/2^n}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   966
  let ?g = "\<lambda>n x. (\<Sum>k::real | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). k/2^n * ?\<Omega> n k x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   967
  have "\<exists>g. (\<forall>n x. 0 \<le> g n x \<and> g n x \<le> (g(Suc n) x)) \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   968
             (\<forall>n. g n \<in> borel_measurable lebesgue) \<and> (\<forall>n. finite(range (g n))) \<and>(\<forall>x. (\<lambda>n. g n x) \<longlonglongrightarrow> f x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   969
  proof (intro exI allI conjI)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   970
    show "0 \<le> ?g n x" for n x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   971
    proof (clarify intro!: ordered_comm_monoid_add_class.sum_nonneg)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   972
      fix k::real
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   973
      assume "k \<in> \<int>" and k: "\<bar>k\<bar> \<le> 2 ^ (2*n)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   974
      show "0 \<le> k/2^n * ?\<Omega> n k x"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   975
        using f \<open>k \<in> \<int>\<close> apply (clarsimp simp: indicator_def field_split_simps Ints_def)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
   976
        by (smt (verit) int_less_real_le mult_nonneg_nonneg of_int_0 zero_le_power)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   977
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   978
    show "?g n x \<le> ?g (Suc n) x" for n x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   979
    proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   980
      have "?g n x =
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   981
            (\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n).
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   982
              k/2^n * (indicator {y. k/2^n \<le> f y \<and> f y < (k+1/2)/2^n} x +
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   983
              indicator {y. (k+1/2)/2^n \<le> f y \<and> f y < (k+1)/2^n} x))"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70802
diff changeset
   984
        by (rule sum.cong [OF refl]) (simp add: indicator_def field_split_simps)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   985
      also have "\<dots> = (\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). k/2^n * indicator {y. k/2^n \<le> f y \<and> f y < (k+1/2)/2^n} x) +
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   986
                       (\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). k/2^n * indicator {y. (k+1/2)/2^n \<le> f y \<and> f y < (k+1)/2^n} x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   987
        by (simp add:  comm_monoid_add_class.sum.distrib algebra_simps)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   988
      also have "\<dots> = (\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). (2 * k)/2 ^ Suc n * indicator {y. (2 * k)/2 ^ Suc n \<le> f y \<and> f y < (2 * k+1)/2 ^ Suc n} x) +
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   989
                       (\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). (2 * k)/2 ^ Suc n * indicator {y. (2 * k+1)/2 ^ Suc n \<le> f y \<and> f y < ((2 * k+1) + 1)/2 ^ Suc n} x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   990
        by (force simp: field_simps indicator_def intro: sum.cong)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   991
      also have "\<dots> \<le> (\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2 * Suc n). k/2 ^ Suc n * (indicator {y. k/2 ^ Suc n \<le> f y \<and> f y < (k+1)/2 ^ Suc n} x))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   992
                (is "?a + _ \<le> ?b")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   993
      proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   994
        have *: "\<lbrakk>sum f I \<le> sum h I; a + sum h I \<le> b\<rbrakk> \<Longrightarrow> a + sum f I \<le> b" for I a b f and h :: "real\<Rightarrow>real"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   995
          by linarith
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   996
        let ?h = "\<lambda>k. (2*k+1)/2 ^ Suc n *
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   997
                      (indicator {y. (2 * k+1)/2 ^ Suc n \<le> f y \<and> f y < ((2*k+1) + 1)/2 ^ Suc n} x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   998
        show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   999
        proof (rule *)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1000
          show "(\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n).
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1001
                  2 * k/2 ^ Suc n * indicator {y. (2 * k+1)/2 ^ Suc n \<le> f y \<and> f y < (2 * k+1 + 1)/2 ^ Suc n} x)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1002
                \<le> sum ?h {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70802
diff changeset
  1003
            by (rule sum_mono) (simp add: indicator_def field_split_simps)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1004
        next
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
  1005
          have \<alpha>: "?a = (\<Sum>k \<in> (*)2 ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}.
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1006
                         k/2 ^ Suc n * indicator {y. k/2 ^ Suc n \<le> f y \<and> f y < (k+1)/2 ^ Suc n} x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1007
            by (auto simp: inj_on_def field_simps comm_monoid_add_class.sum.reindex)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1008
          have \<beta>: "sum ?h {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1009
                   = (\<Sum>k \<in> (\<lambda>x. 2*x + 1) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}.
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1010
                      k/2 ^ Suc n * indicator {y. k/2 ^ Suc n \<le> f y \<and> f y < (k+1)/2 ^ Suc n} x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1011
            by (auto simp: inj_on_def field_simps comm_monoid_add_class.sum.reindex)
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
  1012
          have 0: "(*) 2 ` {k \<in> \<int>. P k} \<inter> (\<lambda>x. 2 * x + 1) ` {k \<in> \<int>. P k} = {}" for P :: "real \<Rightarrow> bool"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1013
          proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1014
            have "2 * i \<noteq> 2 * j + 1" for i j :: int by arith
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1015
            thus ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1016
              unfolding Ints_def by auto (use of_int_eq_iff in fastforce)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1017
          qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1018
          have "?a + sum ?h {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
  1019
                = (\<Sum>k \<in> (*)2 ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)} \<union> (\<lambda>x. 2*x + 1) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}.
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1020
                  k/2 ^ Suc n * indicator {y. k/2 ^ Suc n \<le> f y \<and> f y < (k+1)/2 ^ Suc n} x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1021
            unfolding \<alpha> \<beta>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1022
            using finite_abs_int_segment [of "2 ^ (2*n)"]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1023
            by (subst sum_Un) (auto simp: 0)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1024
          also have "\<dots> \<le> ?b"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1025
          proof (rule sum_mono2)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1026
            show "finite {k::real. k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2 * Suc n)}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1027
              by (rule finite_abs_int_segment)
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
  1028
            show "(*) 2 ` {k::real. k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2^(2*n)} \<union> (\<lambda>x. 2*x + 1) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2^(2*n)} \<subseteq> {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2 * Suc n)}"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  1029
              apply (clarsimp simp: image_subset_iff)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1030
              using one_le_power [of "2::real" "2*n"]  by linarith
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1031
            have *: "\<lbrakk>x \<in> (S \<union> T) - U; \<And>x. x \<in> S \<Longrightarrow> x \<in> U; \<And>x. x \<in> T \<Longrightarrow> x \<in> U\<rbrakk> \<Longrightarrow> P x" for S T U P
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1032
              by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1033
            have "0 \<le> b" if "b \<in> \<int>" "f x * (2 * 2^n) < b + 1" for b
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  1034
              by (smt (verit, ccfv_SIG) Ints_cases f int_le_real_less mult_nonneg_nonneg of_int_add one_le_power that)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1035
            then show "0 \<le> b/2 ^ Suc n * indicator {y. b/2 ^ Suc n \<le> f y \<and> f y < (b + 1)/2 ^ Suc n} x"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1036
                  if "b \<in> {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2 * Suc n)} -
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
  1037
                          ((*) 2 ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)} \<union> (\<lambda>x. 2*x + 1) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)})" for b
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1038
              using that by (simp add: indicator_def divide_simps)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1039
          qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1040
          finally show "?a + sum ?h {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)} \<le> ?b" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1041
        qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1042
      qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1043
      finally show ?thesis .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1044
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1045
    show "?g n \<in> borel_measurable lebesgue" for n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1046
      apply (intro borel_measurable_indicator borel_measurable_times borel_measurable_sum)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1047
      using leb_f sets_restrict_UNIV by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1048
    show "finite (range (?g n))" for n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1049
    proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1050
      have "(\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). k/2^n * ?\<Omega> n k x)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1051
              \<in> (\<lambda>k. k/2^n) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)}" for x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1052
      proof (cases "\<exists>k. k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n) \<and> k/2^n \<le> f x \<and> f x < (k+1)/2^n")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1053
        case True
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1054
        then show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1055
          by (blast intro: indicator_sum_eq)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1056
      next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1057
        case False
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1058
        then have "(\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). k/2^n * ?\<Omega> n k x) = 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1059
          by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1060
        then show ?thesis by force
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1061
      qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1062
      then have "range (?g n) \<subseteq> ((\<lambda>k. (k/2^n)) ` {k. k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n)})"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1063
        by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1064
      moreover have "finite ((\<lambda>k::real. (k/2^n)) ` {k \<in> \<int>. \<bar>k\<bar> \<le> 2 ^ (2*n)})"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1065
        by (intro finite_imageI finite_abs_int_segment)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1066
      ultimately show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1067
        by (rule finite_subset)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1068
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1069
    show "(\<lambda>n. ?g n x) \<longlonglongrightarrow> f x" for x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1070
    proof (clarsimp simp add: lim_sequentially)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1071
      fix e::real
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1072
      assume "e > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1073
      obtain N1 where N1: "2 ^ N1 > abs(f x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1074
        using real_arch_pow by fastforce
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1075
      obtain N2 where N2: "(1/2) ^ N2 < e"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1076
        using real_arch_pow_inv \<open>e > 0\<close> by fastforce
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1077
      have "dist (\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). k/2^n * ?\<Omega> n k x) (f x) < e" if "N1 + N2 \<le> n" for n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1078
      proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1079
        let ?m = "real_of_int \<lfloor>2^n * f x\<rfloor>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1080
        have "\<bar>?m\<bar> \<le> 2^n * 2^N1"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1081
          using N1 apply (simp add: f)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1082
          by (meson floor_mono le_floor_iff less_le_not_le mult_le_cancel_left_pos zero_less_numeral zero_less_power)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1083
        also have "\<dots> \<le> 2 ^ (2*n)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1084
          by (metis that add_leD1 add_le_cancel_left mult.commute mult_2_right one_less_numeral_iff
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1085
                    power_add power_increasing_iff semiring_norm(76))
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1086
        finally have m_le: "\<bar>?m\<bar> \<le> 2 ^ (2*n)" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1087
        have "?m/2^n \<le> f x" "f x < (?m + 1)/2^n"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1088
          by (auto simp: mult.commute pos_divide_le_eq mult_imp_less_div_pos)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1089
        then have eq: "dist (\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). k/2^n * ?\<Omega> n k x) (f x)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1090
                     = dist (?m/2^n) (f x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1091
          by (subst indicator_sum_eq [of ?m]) (auto simp: m_le)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1092
        have "\<bar>2^n\<bar> * \<bar>?m/2^n - f x\<bar> = \<bar>2^n * (?m/2^n - f x)\<bar>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1093
          by (simp add: abs_mult)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1094
        also have "\<dots> < 2 ^ N2 * e"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1095
          using N2 by (simp add: divide_simps mult.commute) linarith
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1096
        also have "\<dots> \<le> \<bar>2^n\<bar> * e"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1097
          using that \<open>e > 0\<close> by auto
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  1098
        finally show ?thesis
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  1099
          using eq by (simp add: dist_real_def)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1100
      qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1101
      then show "\<exists>no. \<forall>n\<ge>no. dist (\<Sum>k | k \<in> \<int> \<and> \<bar>k\<bar> \<le> 2 ^ (2*n). k * ?\<Omega> n k x/2^n) (f x) < e"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1102
        by force
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1103
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1104
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1105
  ultimately show ?rhs
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1106
    by metis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1107
next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1108
  assume RHS: ?rhs
70707
125705f5965f A little-known material, and some tidying up
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
  1109
  with borel_measurable_simple_function_limit [of f UNIV, unfolded lebesgue_on_UNIV_eq]
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1110
  show ?lhs
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1111
    by (blast intro: order_trans)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1112
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1113
69683
8b3458ca0762 subsection is always %important
immler
parents: 69679
diff changeset
  1114
subsection\<open>Borel measurable Jacobian determinant\<close>
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1115
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1116
lemma lemma_partial_derivatives0:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1117
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1118
  assumes "linear f" and lim0: "((\<lambda>x. f x /\<^sub>R norm x) \<longlongrightarrow> 0) (at 0 within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1119
    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>)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1120
  shows "f x = 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1121
proof -
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 68001
diff changeset
  1122
  interpret linear f by fact
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1123
  have "dim {x. f x = 0} \<le> DIM('a)"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 68001
diff changeset
  1124
    by (rule dim_subset_UNIV)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1125
  moreover have False if less: "dim {x. f x = 0} < DIM('a)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1126
  proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1127
    obtain d where "d \<noteq> 0" and d: "\<And>y. f y = 0 \<Longrightarrow> d \<bullet> y = 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1128
      using orthogonal_to_subspace_exists [OF less] orthogonal_def
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
  1129
      by (metis (mono_tags, lifting) mem_Collect_eq span_base)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1130
    then obtain k where "k > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1131
      and k: "\<And>e. e > 0 \<Longrightarrow> \<exists>y. y \<in> S - {0} \<and> norm y < e \<and> k * norm y \<le> \<bar>d \<bullet> y\<bar>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1132
      using lb by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1133
    have "\<exists>h. \<forall>n. ((h n \<in> S \<and> h n \<noteq> 0 \<and> k * norm (h n) \<le> \<bar>d \<bullet> h n\<bar>) \<and> norm (h n) < 1 / real (Suc n)) \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1134
               norm (h (Suc n)) < norm (h n)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1135
    proof (rule dependent_nat_choice)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1136
      show "\<exists>y. (y \<in> S \<and> y \<noteq> 0 \<and> k * norm y \<le> \<bar>d \<bullet> y\<bar>) \<and> norm y < 1 / real (Suc 0)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1137
        by simp (metis DiffE insertCI k not_less not_one_le_zero)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1138
    qed (use k [of "min (norm x) (1/(Suc n + 1))" for x n] in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1139
    then obtain \<alpha> where \<alpha>: "\<And>n. \<alpha> n \<in> S - {0}" and kd: "\<And>n. k * norm(\<alpha> n) \<le> \<bar>d \<bullet> \<alpha> n\<bar>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1140
         and norm_lt: "\<And>n. norm(\<alpha> n) < 1/(Suc n)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1141
      by force
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1142
    let ?\<beta> = "\<lambda>n. \<alpha> n /\<^sub>R norm (\<alpha> n)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1143
    have com: "\<And>g. (\<forall>n. g n \<in> sphere (0::'a) 1)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1144
              \<Longrightarrow> \<exists>l \<in> sphere 0 1. \<exists>\<rho>::nat\<Rightarrow>nat. strict_mono \<rho> \<and> (g \<circ> \<rho>) \<longlonglongrightarrow> l"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1145
      using compact_sphere compact_def by metis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1146
    moreover have "\<forall>n. ?\<beta> n \<in> sphere 0 1"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1147
      using \<alpha> by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1148
    ultimately obtain l::'a and \<rho>::"nat\<Rightarrow>nat"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1149
       where l: "l \<in> sphere 0 1" and "strict_mono \<rho>" and to_l: "(?\<beta> \<circ> \<rho>) \<longlonglongrightarrow> l"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1150
      by meson
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1151
    moreover have "continuous (at l) (\<lambda>x. (\<bar>d \<bullet> x\<bar> - k))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1152
      by (intro continuous_intros)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1153
    ultimately have lim_dl: "((\<lambda>x. (\<bar>d \<bullet> x\<bar> - k)) \<circ> (?\<beta> \<circ> \<rho>)) \<longlonglongrightarrow> (\<bar>d \<bullet> l\<bar> - k)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1154
      by (meson continuous_imp_tendsto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1155
    have "\<forall>\<^sub>F i in sequentially. 0 \<le> ((\<lambda>x. \<bar>d \<bullet> x\<bar> - k) \<circ> ((\<lambda>n. \<alpha> n /\<^sub>R norm (\<alpha> n)) \<circ> \<rho>)) i"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70802
diff changeset
  1156
      using \<alpha> kd by (auto simp: field_split_simps)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1157
    then have "k \<le> \<bar>d \<bullet> l\<bar>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1158
      using tendsto_lowerbound [OF lim_dl, of 0] by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1159
    moreover have "d \<bullet> l = 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1160
    proof (rule d)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1161
      show "f l = 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1162
      proof (rule LIMSEQ_unique [of "f \<circ> ?\<beta> \<circ> \<rho>"])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1163
        have "isCont f l"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1164
          using \<open>linear f\<close> linear_continuous_at linear_conv_bounded_linear by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1165
        then show "(f \<circ> (\<lambda>n. \<alpha> n /\<^sub>R norm (\<alpha> n)) \<circ> \<rho>) \<longlonglongrightarrow> f l"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1166
          unfolding comp_assoc
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1167
          using to_l continuous_imp_tendsto by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1168
        have "\<alpha> \<longlonglongrightarrow> 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1169
          using norm_lt LIMSEQ_norm_0 by metis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1170
        with \<open>strict_mono \<rho>\<close> have "(\<alpha> \<circ> \<rho>) \<longlonglongrightarrow> 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1171
          by (metis LIMSEQ_subseq_LIMSEQ)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1172
        with lim0 \<alpha> have "((\<lambda>x. f x /\<^sub>R norm x) \<circ> (\<alpha> \<circ> \<rho>)) \<longlonglongrightarrow> 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1173
          by (force simp: tendsto_at_iff_sequentially)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1174
        then show "(f \<circ> (\<lambda>n. \<alpha> n /\<^sub>R norm (\<alpha> n)) \<circ> \<rho>) \<longlonglongrightarrow> 0"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 68001
diff changeset
  1175
          by (simp add: o_def scale)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1176
      qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1177
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1178
    ultimately show False
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1179
      using \<open>k > 0\<close> by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1180
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1181
  ultimately have dim: "dim {x. f x = 0} = DIM('a)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1182
    by force
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1183
  then show ?thesis
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  1184
    by (metis (mono_tags, lifting) dim_eq_full UNIV_I eq_0_on_span mem_Collect_eq span_raw_def)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1185
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1186
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1187
lemma lemma_partial_derivatives:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1188
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1189
  assumes "linear f" and lim: "((\<lambda>x. f (x - a) /\<^sub>R norm (x - a)) \<longlongrightarrow> 0) (at a within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1190
    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>)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1191
  shows "f x = 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1192
proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1193
  have "((\<lambda>x. f x /\<^sub>R norm x) \<longlongrightarrow> 0) (at 0 within (\<lambda>x. x-a) ` S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1194
    using lim by (simp add: Lim_within dist_norm)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1195
  then show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1196
  proof (rule lemma_partial_derivatives0 [OF \<open>linear f\<close>])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1197
    fix v :: "'a"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1198
    assume v: "v \<noteq> 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1199
    show "\<exists>k>0. \<forall>e>0. \<exists>x. x \<in> (\<lambda>x. x - a) ` S - {0} \<and> norm x < e \<and> k * norm x \<le> \<bar>v \<bullet> x\<bar>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1200
      using lb [OF v] by (force simp:  norm_minus_commute)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1201
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1202
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1203
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1204
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1205
proposition borel_measurable_partial_derivatives:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1206
  fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1207
  assumes S: "S \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1208
    and f: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1209
  shows "(\<lambda>x. (matrix(f' x)$m$n)) \<in> borel_measurable (lebesgue_on S)"
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1210
proof -
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1211
  have contf: "continuous_on S f"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1212
    using continuous_on_eq_continuous_within f has_derivative_continuous by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1213
  have "{x \<in> S.  (matrix (f' x)$m$n) \<le> b} \<in> sets lebesgue" for b
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1214
  proof (rule sets_negligible_symdiff)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1215
    let ?T = "{x \<in> S. \<forall>e>0. \<exists>d>0. \<exists>A. A$m$n < b \<and> (\<forall>i j. A$i$j \<in> \<rat>) \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1216
                       (\<forall>y \<in> S. norm(y - x) < d \<longrightarrow> norm(f y - f x - A *v (y - x)) \<le> e * norm(y - x))}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1217
    let ?U = "S \<inter>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1218
              (\<Inter>e \<in> {e \<in> \<rat>. e > 0}.
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1219
                \<Union>A \<in> {A. A$m$n < b \<and> (\<forall>i j. A$i$j \<in> \<rat>)}.
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1220
                  \<Union>d \<in> {d \<in> \<rat>. 0 < d}.
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1221
                     S \<inter> (\<Inter>y \<in> S. {x \<in> S. norm(y - x) < d \<longrightarrow> norm(f y - f x - A *v (y - x)) \<le> e * norm(y - x)}))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1222
    have "?T = ?U"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1223
    proof (intro set_eqI iffI)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1224
      fix x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1225
      assume xT: "x \<in> ?T"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1226
      then show "x \<in> ?U"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1227
      proof (clarsimp simp add:)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1228
        fix q :: real
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1229
        assume "q \<in> \<rat>" "q > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1230
        then obtain d A where "d > 0" and A: "A $ m $ n < b" "\<And>i j. A $ i $ j \<in> \<rat>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1231
          "\<And>y. \<lbrakk>y\<in>S;  norm (y - x) < d\<rbrakk> \<Longrightarrow> norm (f y - f x - A *v (y - x)) \<le> q * norm (y - x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1232
          using xT by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1233
        then obtain \<delta> where "d > \<delta>" "\<delta> > 0" "\<delta> \<in> \<rat>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1234
          using Rats_dense_in_real by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1235
        with A show "\<exists>A. A $ m $ n < b \<and> (\<forall>i j. A $ i $ j \<in> \<rat>) \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1236
                         (\<exists>s. s \<in> \<rat> \<and> 0 < s \<and> (\<forall>y\<in>S. norm (y - x) < s \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> q * norm (y - x)))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1237
          by force
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1238
      qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1239
    next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1240
      fix x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1241
      assume xU: "x \<in> ?U"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1242
      then show "x \<in> ?T"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1243
      proof clarsimp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1244
        fix e :: "real"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1245
        assume "e > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1246
        then obtain \<epsilon> where \<epsilon>: "e > \<epsilon>" "\<epsilon> > 0" "\<epsilon> \<in> \<rat>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1247
          using Rats_dense_in_real by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1248
        with xU obtain A r where "x \<in> S" and Ar: "A $ m $ n < b" "\<forall>i j. A $ i $ j \<in> \<rat>" "r \<in> \<rat>" "r > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1249
          and "\<forall>y\<in>S. norm (y - x) < r \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> \<epsilon> * norm (y - x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1250
          by (auto simp: split: if_split_asm)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1251
        then have "\<forall>y\<in>S. norm (y - x) < r \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> e * norm (y - x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1252
          by (meson \<open>e > \<epsilon>\<close> less_eq_real_def mult_right_mono norm_ge_zero order_trans)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1253
        then show "\<exists>d>0. \<exists>A. A $ m $ n < b \<and> (\<forall>i j. A $ i $ j \<in> \<rat>) \<and> (\<forall>y\<in>S. norm (y - x) < d \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> e * norm (y - x))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1254
          using \<open>x \<in> S\<close> Ar by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1255
      qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1256
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1257
    moreover have "?U \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1258
    proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1259
      have coQ: "countable {e \<in> \<rat>. 0 < e}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1260
        using countable_Collect countable_rat by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1261
      have ne: "{e \<in> \<rat>. (0::real) < e} \<noteq> {}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1262
        using zero_less_one Rats_1 by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1263
      have coA: "countable {A. A $ m $ n < b \<and> (\<forall>i j. A $ i $ j \<in> \<rat>)}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1264
      proof (rule countable_subset)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1265
        show "countable {A. \<forall>i j. A $ i $ j \<in> \<rat>}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1266
          using countable_vector [OF countable_vector, of "\<lambda>i j. \<rat>"] by (simp add: countable_rat)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1267
      qed blast
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69739
diff changeset
  1268
      have *: "\<lbrakk>U \<noteq> {} \<Longrightarrow> closedin (top_of_set S) (S \<inter> \<Inter> U)\<rbrakk>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69739
diff changeset
  1269
               \<Longrightarrow> closedin (top_of_set S) (S \<inter> \<Inter> U)" for U
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1270
        by fastforce
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1271
      have eq: "{x::(real,'m)vec. P x \<and> (Q x \<longrightarrow> R x)} = {x. P x \<and> \<not> Q x} \<union> {x. P x \<and> R x}" for P Q R
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1272
        by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1273
      have sets: "S \<inter> (\<Inter>y\<in>S. {x \<in> S. norm (y - x) < d \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> e * norm (y - x)})
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1274
                  \<in> sets lebesgue" for e A d
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1275
      proof -
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69739
diff changeset
  1276
        have clo: "closedin (top_of_set S)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1277
                     {x \<in> S. norm (y - x) < d \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> e * norm (y - x)}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1278
          for y
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1279
        proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1280
          have cont1: "continuous_on S (\<lambda>x. norm (y - x))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1281
          and  cont2: "continuous_on S (\<lambda>x. e * norm (y - x) - norm (f y - f x - (A *v y - A *v x)))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1282
            by (force intro: contf continuous_intros)+
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69739
diff changeset
  1283
          have clo1: "closedin (top_of_set S) {x \<in> S. d \<le> norm (y - x)}"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1284
            using continuous_closedin_preimage [OF cont1, of "{d..}"] by (simp add: vimage_def Int_def)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69739
diff changeset
  1285
          have clo2: "closedin (top_of_set S)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1286
                       {x \<in> S. norm (f y - f x - (A *v y - A *v x)) \<le> e * norm (y - x)}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1287
            using continuous_closedin_preimage [OF cont2, of "{0..}"] by (simp add: vimage_def Int_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1288
          show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1289
            by (auto simp: eq not_less matrix_vector_mult_diff_distrib intro: clo1 clo2)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1290
        qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1291
        show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1292
          by (rule lebesgue_closedin [of S]) (force intro: * S clo)+
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1293
      qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1294
      show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1295
        by (intro sets sets.Int S sets.countable_UN'' sets.countable_INT'' coQ coA) auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1296
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1297
    ultimately show "?T \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1298
      by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1299
    let ?M = "(?T - {x \<in> S. matrix (f' x) $ m $ n \<le> b} \<union> ({x \<in> S. matrix (f' x) $ m $ n \<le> b} - ?T))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1300
    let ?\<Theta> = "\<lambda>x v. \<forall>\<xi>>0. \<exists>e>0. \<forall>y \<in> S-{x}. norm (x - y) < e \<longrightarrow> \<bar>v \<bullet> (y - x)\<bar> < \<xi> * norm (x - y)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1301
    have nN: "negligible {x \<in> S. \<exists>v\<noteq>0. ?\<Theta> x v}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1302
      unfolding negligible_eq_zero_density
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1303
    proof clarsimp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1304
      fix x v and r e :: "real"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1305
      assume "x \<in> S" "v \<noteq> 0" "r > 0" "e > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1306
      and Theta [rule_format]: "?\<Theta> x v"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1307
      moreover have "(norm v * e / 2) / CARD('m) ^ CARD('m) > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1308
        by (simp add: \<open>v \<noteq> 0\<close> \<open>e > 0\<close>)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1309
      ultimately obtain d where "d > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1310
         and dless: "\<And>y. \<lbrakk>y \<in> S - {x}; norm (x - y) < d\<rbrakk> \<Longrightarrow>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1311
                        \<bar>v \<bullet> (y - x)\<bar> < ((norm v * e / 2) / CARD('m) ^ CARD('m)) * norm (x - y)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1312
        by metis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1313
      let ?W = "ball x (min d r) \<inter> {y. \<bar>v \<bullet> (y - x)\<bar> < (norm v * e/2 * min d r) / CARD('m) ^ CARD('m)}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1314
      have "open {x. \<bar>v \<bullet> (x - a)\<bar> < b}" for a b
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1315
        by (intro open_Collect_less continuous_intros)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1316
      show "\<exists>d>0. d \<le> r \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1317
            (\<exists>U. {x' \<in> S. \<exists>v\<noteq>0. ?\<Theta> x' v} \<inter> ball x d \<subseteq> U \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1318
                 U \<in> lmeasurable \<and> measure lebesgue U < e * content (ball x d))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1319
      proof (intro exI conjI)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1320
        show "0 < min d r" "min d r \<le> r"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1321
          using \<open>r > 0\<close> \<open>d > 0\<close> by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1322
        show "{x' \<in> S. \<exists>v. v \<noteq> 0 \<and> (\<forall>\<xi>>0. \<exists>e>0. \<forall>z\<in>S - {x'}. norm (x' - z) < e \<longrightarrow> \<bar>v \<bullet> (z - x')\<bar> < \<xi> * norm (x' - z))} \<inter> ball x (min d r) \<subseteq> ?W"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1323
          proof (clarsimp simp: dist_norm norm_minus_commute)
68001
0a2a1b6507c1 correction of TeX errors and other oversights
paulson <lp15@cam.ac.uk>
parents: 67999
diff changeset
  1324
            fix y w
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1325
            assume "y \<in> S" "w \<noteq> 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1326
              and less [rule_format]:
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1327
                    "\<forall>\<xi>>0. \<exists>e>0. \<forall>z\<in>S - {y}. norm (y - z) < e \<longrightarrow> \<bar>w \<bullet> (z - y)\<bar> < \<xi> * norm (y - z)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1328
              and d: "norm (y - x) < d" and r: "norm (y - x) < r"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1329
            show "\<bar>v \<bullet> (y - x)\<bar> < norm v * e * min d r / (2 * real CARD('m) ^ CARD('m))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1330
            proof (cases "y = x")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1331
              case True
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1332
              with \<open>r > 0\<close> \<open>d > 0\<close> \<open>e > 0\<close> \<open>v \<noteq> 0\<close> show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1333
                by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1334
            next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1335
              case False
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1336
              have "\<bar>v \<bullet> (y - x)\<bar> < norm v * e / 2 / real (CARD('m) ^ CARD('m)) * norm (x - y)"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  1337
                by (metis Diff_iff False \<open>y \<in> S\<close> d dless empty_iff insert_iff norm_minus_commute)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1338
              also have "\<dots> \<le> norm v * e * min d r / (2 * real CARD('m) ^ CARD('m))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1339
                using d r \<open>e > 0\<close> by (simp add: field_simps norm_minus_commute mult_left_mono)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1340
              finally show ?thesis .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1341
            qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1342
          qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1343
          show "?W \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1344
            by (simp add: fmeasurable_Int_fmeasurable borel_open)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1345
          obtain k::'m where True
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1346
            by metis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1347
          obtain T where T: "orthogonal_transformation T" and v: "v = T(norm v *\<^sub>R axis k (1::real))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1348
            using rotation_rightward_line by metis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1349
          define b where "b \<equiv> norm v"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1350
          have "b > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1351
            using \<open>v \<noteq> 0\<close> by (auto simp: b_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1352
          obtain eqb: "inv T v = b *\<^sub>R axis k (1::real)" and "inj T" "bij T" and invT: "orthogonal_transformation (inv T)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1353
            by (metis UNIV_I b_def  T v bij_betw_inv_into_left orthogonal_transformation_inj orthogonal_transformation_bij orthogonal_transformation_inv)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1354
          let ?v = "\<chi> i. min d r / CARD('m)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1355
          let ?v' = "\<chi> i. if i = k then (e/2 * min d r) / CARD('m) ^ CARD('m) else min d r"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1356
          let ?x' = "inv T x"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1357
          let ?W' = "(ball ?x' (min d r) \<inter> {y. \<bar>(y - ?x')$k\<bar> < e * min d r / (2 * CARD('m) ^ CARD('m))})"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1358
          have abs: "x - e \<le> y \<and> y \<le> x + e \<longleftrightarrow> abs(y - x) \<le> e" for x y e::real
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1359
            by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1360
          have "?W = T ` ?W'"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1361
          proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1362
            have 1: "T ` (ball (inv T x) (min d r)) = ball x (min d r)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1363
              by (simp add: T image_orthogonal_transformation_ball orthogonal_transformation_surj surj_f_inv_f)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1364
            have 2: "{y. \<bar>v \<bullet> (y - x)\<bar> < b * e * min d r / (2 * real CARD('m) ^ CARD('m))} =
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1365
                      T ` {y. \<bar>y $ k - ?x' $ k\<bar> < e * min d r / (2 * real CARD('m) ^ CARD('m))}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1366
            proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1367
              have *: "\<bar>T (b *\<^sub>R axis k 1) \<bullet> (y - x)\<bar> = b * \<bar>inv T y $ k - ?x' $ k\<bar>" for y
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1368
              proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1369
                have "\<bar>T (b *\<^sub>R axis k 1) \<bullet> (y - x)\<bar> = \<bar>(b *\<^sub>R axis k 1) \<bullet> inv T (y - x)\<bar>"
73932
fd21b4a93043 added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents: 73648
diff changeset
  1370
                  by (metis (no_types, opaque_lifting) b_def eqb invT orthogonal_transformation_def v)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1371
                also have "\<dots> = b * \<bar>(axis k 1) \<bullet> inv T (y - x)\<bar>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1372
                  using \<open>b > 0\<close> by (simp add: abs_mult)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1373
                also have "\<dots> = b * \<bar>inv T y $ k - ?x' $ k\<bar>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1374
                  using orthogonal_transformation_linear [OF invT]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1375
                  by (simp add: inner_axis' linear_diff)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1376
                finally show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1377
                  by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1378
              qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1379
              show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1380
                using v b_def [symmetric]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1381
                using \<open>b > 0\<close> by (simp add: * bij_image_Collect_eq [OF \<open>bij T\<close>] mult_less_cancel_left_pos times_divide_eq_right [symmetric] del: times_divide_eq_right)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1382
            qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1383
            show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1384
              using \<open>b > 0\<close> by (simp add: image_Int \<open>inj T\<close> 1 2 b_def [symmetric])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1385
          qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1386
          moreover have "?W' \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1387
            by (auto intro: fmeasurable_Int_fmeasurable)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1388
          ultimately have "measure lebesgue ?W = measure lebesgue ?W'"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1389
            by (metis measure_orthogonal_image T)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1390
          also have "\<dots> \<le> measure lebesgue (cbox (?x' - ?v') (?x' + ?v'))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1391
          proof (rule measure_mono_fmeasurable)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1392
            show "?W' \<subseteq> cbox (?x' - ?v') (?x' + ?v')"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1393
              apply (clarsimp simp add: mem_box_cart abs dist_norm norm_minus_commute simp del: min_less_iff_conj min.bounded_iff)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1394
              by (metis component_le_norm_cart less_eq_real_def le_less_trans vector_minus_component)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1395
          qed auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1396
          also have "\<dots> \<le> e/2 * measure lebesgue (cbox (?x' - ?v) (?x' + ?v))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1397
          proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1398
            have "cbox (?x' - ?v) (?x' + ?v) \<noteq> {}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1399
              using \<open>r > 0\<close> \<open>d > 0\<close> by (auto simp: interval_eq_empty_cart divide_less_0_iff)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1400
            with \<open>r > 0\<close> \<open>d > 0\<close> \<open>e > 0\<close> show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1401
              apply (simp add: content_cbox_if_cart mem_box_cart)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1402
              apply (auto simp: prod_nonneg)
71172
nipkow
parents: 70817
diff changeset
  1403
              apply (simp add: abs if_distrib prod.delta_remove field_simps power_diff split: if_split_asm)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1404
              done
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1405
          qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1406
          also have "\<dots> \<le> e/2 * measure lebesgue (cball ?x' (min d r))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1407
          proof (rule mult_left_mono [OF measure_mono_fmeasurable])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1408
            have *: "norm (?x' - y) \<le> min d r"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1409
              if y: "\<And>i. \<bar>?x' $ i - y $ i\<bar> \<le> min d r / real CARD('m)" for y
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1410
            proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1411
              have "norm (?x' - y) \<le> (\<Sum>i\<in>UNIV. \<bar>(?x' - y) $ i\<bar>)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1412
                by (rule norm_le_l1_cart)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1413
              also have "\<dots> \<le> real CARD('m) * (min d r / real CARD('m))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1414
                by (rule sum_bounded_above) (use y in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1415
              finally show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1416
                by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1417
            qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1418
            show "cbox (?x' - ?v) (?x' + ?v) \<subseteq> cball ?x' (min d r)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1419
              apply (clarsimp simp only: mem_box_cart dist_norm mem_cball intro!: *)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1420
              by (simp add: abs_diff_le_iff abs_minus_commute)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1421
          qed (use \<open>e > 0\<close> in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1422
          also have "\<dots> < e * content (cball ?x' (min d r))"
71192
a8ccea88b725 Flattened dependency tree of HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71189
diff changeset
  1423
            using \<open>r > 0\<close> \<open>d > 0\<close> \<open>e > 0\<close> by (auto intro: content_cball_pos)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1424
          also have "\<dots> = e * content (ball x (min d r))"
71192
a8ccea88b725 Flattened dependency tree of HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71189
diff changeset
  1425
            using \<open>r > 0\<close> \<open>d > 0\<close> content_ball_conv_unit_ball[of "min d r" "inv T x"]
a8ccea88b725 Flattened dependency tree of HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71189
diff changeset
  1426
                  content_ball_conv_unit_ball[of "min d r" x]
a8ccea88b725 Flattened dependency tree of HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71189
diff changeset
  1427
            by (simp add: content_cball_conv_ball)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1428
          finally show "measure lebesgue ?W < e * content (ball x (min d r))" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1429
      qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1430
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1431
    have *: "(\<And>x. (x \<notin> S) \<Longrightarrow> (x \<in> T \<longleftrightarrow> x \<in> U)) \<Longrightarrow> (T - U) \<union> (U - T) \<subseteq> S" for S T U :: "(real,'m) vec set"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1432
      by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1433
    have MN: "?M \<subseteq> {x \<in> S. \<exists>v\<noteq>0. ?\<Theta> x v}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1434
    proof (rule *)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1435
      fix x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1436
      assume x: "x \<notin> {x \<in> S. \<exists>v\<noteq>0. ?\<Theta> x v}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1437
      show "(x \<in> ?T) \<longleftrightarrow> (x \<in> {x \<in> S. matrix (f' x) $ m $ n \<le> b})"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1438
      proof (cases "x \<in> S")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1439
        case True
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1440
        then have x: "\<not> ?\<Theta> x v" if "v \<noteq> 0" for v
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1441
          using x that by force
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1442
        show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1443
        proof (rule iffI; clarsimp)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1444
          assume b: "\<forall>e>0. \<exists>d>0. \<exists>A. A $ m $ n < b \<and> (\<forall>i j. A $ i $ j \<in> \<rat>) \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1445
                                    (\<forall>y\<in>S. norm (y - x) < d \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> e * norm (y - x))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1446
                     (is "\<forall>e>0. \<exists>d>0. \<exists>A. ?\<Phi> e d A")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1447
          then have "\<forall>k. \<exists>d>0. \<exists>A. ?\<Phi> (1 / Suc k) d A"
73932
fd21b4a93043 added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents: 73648
diff changeset
  1448
            by (metis (no_types, opaque_lifting) less_Suc_eq_0_disj of_nat_0_less_iff zero_less_divide_1_iff)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1449
          then obtain \<delta> A where \<delta>: "\<And>k. \<delta> k > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1450
                           and Ab: "\<And>k. A k $ m $ n < b"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1451
                           and A: "\<And>k y. \<lbrakk>y \<in> S; norm (y - x) < \<delta> k\<rbrakk> \<Longrightarrow>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1452
                                          norm (f y - f x - A k *v (y - x)) \<le> 1/(Suc k) * norm (y - x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1453
            by metis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1454
          have "\<forall>i j. \<exists>a. (\<lambda>n. A n $ i $ j) \<longlonglongrightarrow> a"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1455
          proof (intro allI)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1456
            fix i j
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1457
            have vax: "(A n *v axis j 1) $ i = A n $ i $ j" for n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1458
              by (metis cart_eq_inner_axis matrix_vector_mul_component)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1459
            let ?CA = "{x. Cauchy (\<lambda>n. (A n) *v x)}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1460
            have "subspace ?CA"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1461
              unfolding subspace_def convergent_eq_Cauchy [symmetric]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1462
                by (force simp: algebra_simps intro: tendsto_intros)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1463
            then have CA_eq: "?CA = span ?CA"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 68001
diff changeset
  1464
              by (metis span_eq_iff)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1465
            also have "\<dots> = UNIV"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1466
            proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1467
              have "dim ?CA \<le> CARD('m)"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  1468
                using dim_subset_UNIV[of ?CA] by auto
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1469
              moreover have "False" if less: "dim ?CA < CARD('m)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1470
              proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1471
                obtain d where "d \<noteq> 0" and d: "\<And>y. y \<in> span ?CA \<Longrightarrow> orthogonal d y"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1472
                  using less by (force intro: orthogonal_to_subspace_exists [of ?CA])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1473
                with x [OF \<open>d \<noteq> 0\<close>] obtain \<xi> where "\<xi> > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1474
                  and \<xi>: "\<And>e. e > 0 \<Longrightarrow> \<exists>y \<in> S - {x}. norm (x - y) < e \<and> \<xi> * norm (x - y) \<le> \<bar>d \<bullet> (y - x)\<bar>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1475
                  by (fastforce simp: not_le Bex_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1476
                obtain \<gamma> z where \<gamma>Sx: "\<And>i. \<gamma> i \<in> S - {x}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1477
                           and \<gamma>le:   "\<And>i. \<xi> * norm(\<gamma> i - x) \<le> \<bar>d \<bullet> (\<gamma> i - x)\<bar>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1478
                           and \<gamma>x:    "\<gamma> \<longlonglongrightarrow> x"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1479
                           and z:     "(\<lambda>n. (\<gamma> n - x) /\<^sub>R norm (\<gamma> n - x)) \<longlonglongrightarrow> z"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1480
                proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1481
                  have "\<exists>\<gamma>. (\<forall>i. (\<gamma> i \<in> S - {x} \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1482
                                  \<xi> * norm(\<gamma> i - x) \<le> \<bar>d \<bullet> (\<gamma> i - x)\<bar> \<and> norm(\<gamma> i - x) < 1/Suc i) \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1483
                                 norm(\<gamma>(Suc i) - x) < norm(\<gamma> i - x))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1484
                  proof (rule dependent_nat_choice)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1485
                    show "\<exists>y. y \<in> S - {x} \<and> \<xi> * norm (y - x) \<le> \<bar>d \<bullet> (y - x)\<bar> \<and> norm (y - x) < 1 / Suc 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1486
                      using \<xi> [of 1] by (auto simp: dist_norm norm_minus_commute)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1487
                  next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1488
                    fix y i
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1489
                    assume "y \<in> S - {x} \<and> \<xi> * norm (y - x) \<le> \<bar>d \<bullet> (y - x)\<bar> \<and> norm (y - x) < 1/Suc i"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1490
                    then have "min (norm(y - x)) (1/((Suc i) + 1)) > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1491
                      by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1492
                    then obtain y' where "y' \<in> S - {x}" and y': "norm (x - y') < min (norm (y - x)) (1/((Suc i) + 1))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1493
                                         "\<xi> * norm (x - y') \<le> \<bar>d \<bullet> (y' - x)\<bar>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1494
                      using \<xi> by metis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1495
                    with \<xi> show "\<exists>y'. (y' \<in> S - {x} \<and> \<xi> * norm (y' - x) \<le> \<bar>d \<bullet> (y' - x)\<bar> \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1496
                              norm (y' - x) < 1/(Suc (Suc i))) \<and> norm (y' - x) < norm (y - x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1497
                      by (auto simp: dist_norm norm_minus_commute)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1498
                  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1499
                  then obtain \<gamma> where
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1500
                        \<gamma>Sx: "\<And>i. \<gamma> i \<in> S - {x}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1501
                        and \<gamma>le: "\<And>i. \<xi> * norm(\<gamma> i - x) \<le> \<bar>d \<bullet> (\<gamma> i - x)\<bar>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1502
                        and \<gamma>conv: "\<And>i. norm(\<gamma> i - x) < 1/(Suc i)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1503
                    by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1504
                  let ?f = "\<lambda>i. (\<gamma> i - x) /\<^sub>R norm (\<gamma> i - x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1505
                  have "?f i \<in> sphere 0 1" for i
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1506
                    using \<gamma>Sx by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1507
                  then obtain l \<rho> where "l \<in> sphere 0 1" "strict_mono \<rho>" and l: "(?f \<circ> \<rho>) \<longlonglongrightarrow> l"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1508
                    using compact_sphere [of "0::(real,'m) vec" 1]  unfolding compact_def by meson
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1509
                  show thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1510
                  proof
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1511
                    show "(\<gamma> \<circ> \<rho>) i \<in> S - {x}" "\<xi> * norm ((\<gamma> \<circ> \<rho>) i - x) \<le> \<bar>d \<bullet> ((\<gamma> \<circ> \<rho>) i - x)\<bar>" for i
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1512
                      using \<gamma>Sx \<gamma>le by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1513
                    have "\<gamma> \<longlonglongrightarrow> x"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1514
                    proof (clarsimp simp add: LIMSEQ_def dist_norm)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1515
                      fix r :: "real"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1516
                      assume "r > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1517
                      with real_arch_invD obtain no where "no \<noteq> 0" "real no > 1/r"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1518
                        by (metis divide_less_0_1_iff not_less_iff_gr_or_eq of_nat_0_eq_iff reals_Archimedean2)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1519
                      with \<gamma>conv show "\<exists>no. \<forall>n\<ge>no. norm (\<gamma> n - x) < r"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1520
                        by (metis \<open>r > 0\<close> add.commute divide_inverse inverse_inverse_eq inverse_less_imp_less less_trans mult.left_neutral nat_le_real_less of_nat_Suc)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1521
                    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1522
                    with \<open>strict_mono \<rho>\<close> show "(\<gamma> \<circ> \<rho>) \<longlonglongrightarrow> x"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1523
                      by (metis LIMSEQ_subseq_LIMSEQ)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1524
                    show "(\<lambda>n. ((\<gamma> \<circ> \<rho>) n - x) /\<^sub>R norm ((\<gamma> \<circ> \<rho>) n - x)) \<longlonglongrightarrow> l"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1525
                      using l by (auto simp: o_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1526
                  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1527
                qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1528
                have "isCont (\<lambda>x. (\<bar>d \<bullet> x\<bar> - \<xi>)) z"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1529
                  by (intro continuous_intros)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1530
                from isCont_tendsto_compose [OF this z]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1531
                have lim: "(\<lambda>y. \<bar>d \<bullet> ((\<gamma> y - x) /\<^sub>R norm (\<gamma> y - x))\<bar> - \<xi>) \<longlonglongrightarrow> \<bar>d \<bullet> z\<bar> - \<xi>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1532
                  by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1533
                moreover have "\<forall>\<^sub>F i in sequentially. 0 \<le> \<bar>d \<bullet> ((\<gamma> i - x) /\<^sub>R norm (\<gamma> i - x))\<bar> - \<xi>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1534
                proof (rule eventuallyI)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1535
                  fix n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1536
                  show "0 \<le> \<bar>d \<bullet> ((\<gamma> n - x) /\<^sub>R norm (\<gamma> n - x))\<bar> - \<xi>"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  1537
                    using \<gamma>le [of n] \<gamma>Sx by (auto simp: abs_mult divide_simps)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1538
                qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1539
                ultimately have "\<xi> \<le> \<bar>d \<bullet> z\<bar>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1540
                  using tendsto_lowerbound [where a=0] by fastforce
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1541
                have "Cauchy (\<lambda>n. (A n) *v z)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1542
                proof (clarsimp simp add: Cauchy_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1543
                  fix \<epsilon> :: "real"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1544
                  assume "0 < \<epsilon>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1545
                  then obtain N::nat where "N > 0" and N: "\<epsilon>/2 > 1/N"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1546
                    by (metis half_gt_zero inverse_eq_divide neq0_conv real_arch_inverse)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1547
                  show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (A m *v z) (A n *v z) < \<epsilon>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1548
                  proof (intro exI allI impI)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1549
                    fix i j
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1550
                    assume ij: "N \<le> i" "N \<le> j"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1551
                    let ?V = "\<lambda>i k. A i *v ((\<gamma> k - x) /\<^sub>R norm (\<gamma> k - x))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1552
                    have "\<forall>\<^sub>F k in sequentially. dist (\<gamma> k) x < min (\<delta> i) (\<delta> j)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1553
                      using \<gamma>x [unfolded tendsto_iff] by (meson min_less_iff_conj \<delta>)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1554
                    then have even: "\<forall>\<^sub>F k in sequentially. norm (?V i k - ?V j k) - 2 / N \<le> 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1555
                    proof (rule eventually_mono, clarsimp)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1556
                      fix p
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1557
                      assume p: "dist (\<gamma> p) x < \<delta> i" "dist (\<gamma> p) x < \<delta> j"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1558
                      let ?C = "\<lambda>k. f (\<gamma> p) - f x - A k *v (\<gamma> p - x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1559
                      have "norm ((A i - A j) *v (\<gamma> p - x)) = norm (?C j - ?C i)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1560
                        by (simp add: algebra_simps)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1561
                      also have "\<dots> \<le> norm (?C j) + norm (?C i)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1562
                        using norm_triangle_ineq4 by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1563
                      also have "\<dots> \<le> 1/(Suc j) * norm (\<gamma> p - x) + 1/(Suc i) * norm (\<gamma> p - x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1564
                        by (metis A Diff_iff \<gamma>Sx dist_norm p add_mono)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1565
                      also have "\<dots> \<le> 1/N * norm (\<gamma> p - x) + 1/N * norm (\<gamma> p - x)"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  1566
                        using ij \<open>N > 0\<close> by (intro add_mono mult_right_mono) (auto simp: field_simps)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1567
                      also have "\<dots> = 2 / N * norm (\<gamma> p - x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1568
                        by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1569
                      finally have no_le: "norm ((A i - A j) *v (\<gamma> p - x)) \<le> 2 / N * norm (\<gamma> p - x)" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1570
                      have "norm (?V i p - ?V j p) =
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1571
                            norm ((A i - A j) *v ((\<gamma> p - x) /\<^sub>R norm (\<gamma> p - x)))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1572
                        by (simp add: algebra_simps)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1573
                      also have "\<dots> = norm ((A i - A j) *v (\<gamma> p - x)) / norm (\<gamma> p - x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1574
                        by (simp add: divide_inverse matrix_vector_mult_scaleR)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1575
                      also have "\<dots> \<le> 2 / N"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70802
diff changeset
  1576
                        using no_le by (auto simp: field_split_simps)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1577
                      finally show "norm (?V i p - ?V j p) \<le> 2 / N" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1578
                    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1579
                    have "isCont (\<lambda>w. (norm(A i *v w - A j *v w) - 2 / N)) z"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1580
                      by (intro continuous_intros)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1581
                    from isCont_tendsto_compose [OF this z]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1582
                    have lim: "(\<lambda>w. norm (A i *v ((\<gamma> w - x) /\<^sub>R norm (\<gamma> w - x)) -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1583
                                    A j *v ((\<gamma> w - x) /\<^sub>R norm (\<gamma> w - x))) - 2 / N)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1584
                               \<longlonglongrightarrow> norm (A i *v z - A j *v z) - 2 / N"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1585
                      by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1586
                    have "dist (A i *v z) (A j *v z) \<le> 2 / N"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1587
                      using tendsto_upperbound [OF lim even] by (auto simp: dist_norm)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1588
                    with N show "dist (A i *v z) (A j *v z) < \<epsilon>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1589
                      by linarith
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1590
                  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1591
                qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1592
                then have "d \<bullet> z = 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1593
                  using CA_eq d orthogonal_def by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1594
                then show False
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1595
                  using \<open>0 < \<xi>\<close> \<open>\<xi> \<le> \<bar>d \<bullet> z\<bar>\<close> by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1596
              qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1597
              ultimately show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1598
                using dim_eq_full by fastforce
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1599
            qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1600
            finally have "?CA = UNIV" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1601
            then have "Cauchy (\<lambda>n. (A n) *v axis j 1)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1602
              by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1603
            then obtain L where "(\<lambda>n. A n *v axis j 1) \<longlonglongrightarrow> L"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1604
              by (auto simp: Cauchy_convergent_iff convergent_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1605
            then have "(\<lambda>x. (A x *v axis j 1) $ i) \<longlonglongrightarrow> L $ i"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1606
              by (rule tendsto_vec_nth)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1607
            then show "\<exists>a. (\<lambda>n. A n $ i $ j) \<longlonglongrightarrow> a"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1608
              by (force simp: vax)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1609
          qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1610
          then obtain B where B: "\<And>i j. (\<lambda>n. A n $ i $ j) \<longlonglongrightarrow> B $ i $ j"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1611
            by (auto simp: lambda_skolem)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1612
          have lin_df: "linear (f' x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1613
               and lim_df: "((\<lambda>y. (1 / norm (y - x)) *\<^sub>R (f y - (f x + f' x (y - x)))) \<longlongrightarrow> 0) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1614
            using \<open>x \<in> S\<close> assms by (auto simp: has_derivative_within linear_linear)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 68001
diff changeset
  1615
          moreover
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 68001
diff changeset
  1616
          interpret linear "f' x" by fact
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 68001
diff changeset
  1617
          have "(matrix (f' x) - B) *v w = 0" for w
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
  1618
          proof (rule lemma_partial_derivatives [of "(*v) (matrix (f' x) - B)"])
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
  1619
            show "linear ((*v) (matrix (f' x) - B))"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1620
              by (rule matrix_vector_mul_linear)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1621
            have "((\<lambda>y. ((f x + f' x (y - x)) - f y) /\<^sub>R norm (y - x)) \<longlongrightarrow> 0) (at x within S)"
71633
07bec530f02e cleaned proofs
nipkow
parents: 71192
diff changeset
  1622
              using tendsto_minus [OF lim_df] by (simp add: field_split_simps)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1623
            then show "((\<lambda>y. (matrix (f' x) - B) *v (y - x) /\<^sub>R norm (y - x)) \<longlongrightarrow> 0) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1624
            proof (rule Lim_transform)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1625
              have "((\<lambda>y. ((f y + B *v x - (f x + B *v y)) /\<^sub>R norm (y - x))) \<longlongrightarrow> 0) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1626
              proof (clarsimp simp add: Lim_within dist_norm)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1627
                fix e :: "real"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1628
                assume "e > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1629
                then obtain q::nat where "q \<noteq> 0" and qe2: "1/q < e/2"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1630
                  by (metis divide_pos_pos inverse_eq_divide real_arch_inverse zero_less_numeral)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1631
                let ?g = "\<lambda>p. sum  (\<lambda>i. sum (\<lambda>j. abs((A p - B)$i$j)) UNIV) UNIV"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1632
                have "(\<lambda>k. onorm (\<lambda>y. (A k - B) *v y)) \<longlonglongrightarrow> 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1633
                proof (rule Lim_null_comparison)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1634
                  show "\<forall>\<^sub>F k in sequentially. norm (onorm (\<lambda>y. (A k - B) *v y)) \<le> ?g k"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1635
                  proof (rule eventually_sequentiallyI)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1636
                    fix k :: "nat"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1637
                    assume "0 \<le> k"
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
  1638
                    have "0 \<le> onorm ((*v) (A k - B))"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 68001
diff changeset
  1639
                      using matrix_vector_mul_bounded_linear
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 68001
diff changeset
  1640
                      by (rule onorm_pos_le)
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
  1641
                    then show "norm (onorm ((*v) (A k - B))) \<le> (\<Sum>i\<in>UNIV. \<Sum>j\<in>UNIV. \<bar>(A k - B) $ i $ j\<bar>)"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1642
                      by (simp add: onorm_le_matrix_component_sum del: vector_minus_component)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1643
                  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1644
                next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1645
                  show "?g \<longlonglongrightarrow> 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1646
                    using B Lim_null tendsto_rabs_zero_iff by (fastforce intro!: tendsto_null_sum)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1647
                qed
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
  1648
                with \<open>e > 0\<close> obtain p where "\<And>n. n \<ge> p \<Longrightarrow> \<bar>onorm ((*v) (A n - B))\<bar> < e/2"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1649
                  unfolding lim_sequentially by (metis diff_zero dist_real_def divide_pos_pos zero_less_numeral)
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
  1650
                then have pqe2: "\<bar>onorm ((*v) (A (p + q) - B))\<bar> < e/2" (*17 [`abs (onorm (\y. A (p + q) ** y - B ** y)) < e / &2`]*)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1651
                  using le_add1 by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1652
                show "\<exists>d>0. \<forall>y\<in>S. y \<noteq> x \<and> norm (y - x) < d \<longrightarrow>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1653
                           inverse (norm (y - x)) * norm (f y + B *v x - (f x + B *v y)) < e"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1654
                proof (intro exI, safe)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1655
                  show "0 < \<delta>(p + q)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1656
                    by (simp add: \<delta>)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1657
                next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1658
                  fix y
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1659
                  assume y: "y \<in> S" "norm (y - x) < \<delta>(p + q)" and "y \<noteq> x"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1660
                  have *: "\<lbrakk>norm(b - c) < e - d; norm(y - x - b) \<le> d\<rbrakk> \<Longrightarrow> norm(y - x - c) < e"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1661
                    for b c d e x and y:: "real^'n"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1662
                    using norm_triangle_ineq2 [of "y - x - c" "y - x - b"] by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1663
                  have "norm (f y - f x - B *v (y - x)) < e * norm (y - x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1664
                  proof (rule *)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1665
                    show "norm (f y - f x - A (p + q) *v (y - x)) \<le> norm (y - x) / (Suc (p + q))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1666
                      using A [OF y] by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1667
                    have "norm (A (p + q) *v (y - x) - B *v (y - x)) \<le> onorm(\<lambda>x. (A(p + q) - B) *v x) * norm(y - x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1668
                      by (metis linear_linear matrix_vector_mul_linear matrix_vector_mult_diff_rdistrib onorm)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1669
                    also have "\<dots> < (e/2) * norm (y - x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1670
                      using \<open>y \<noteq> x\<close> pqe2 by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1671
                    also have "\<dots> \<le> (e - 1 / (Suc (p + q))) * norm (y - x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1672
                    proof (rule mult_right_mono)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1673
                      have "1 / Suc (p + q) \<le> 1 / q"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70802
diff changeset
  1674
                        using \<open>q \<noteq> 0\<close> by (auto simp: field_split_simps)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1675
                      also have "\<dots> < e/2"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1676
                        using qe2 by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1677
                      finally show "e / 2 \<le> e - 1 / real (Suc (p + q))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1678
                        by linarith
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1679
                    qed auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1680
                    finally show "norm (A (p + q) *v (y - x) - B *v (y - x)) < e * norm (y - x) - norm (y - x) / real (Suc (p + q))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1681
                      by (simp add: algebra_simps)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1682
                  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1683
                  then show "inverse (norm (y - x)) * norm (f y + B *v x - (f x + B *v y)) < e"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70802
diff changeset
  1684
                    using \<open>y \<noteq> x\<close> by (simp add: field_split_simps algebra_simps)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1685
                qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1686
              qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1687
              then show "((\<lambda>y. (matrix (f' x) - B) *v (y - x) /\<^sub>R
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1688
                           norm (y - x) - (f x + f' x (y - x) - f y) /\<^sub>R norm (y - x)) \<longlongrightarrow> 0)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1689
                          (at x within S)"
71633
07bec530f02e cleaned proofs
nipkow
parents: 71192
diff changeset
  1690
                by (simp add: algebra_simps diff lin_df scalar_mult_eq_scaleR)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1691
            qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1692
          qed (use x in \<open>simp; auto simp: not_less\<close>)
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
  1693
          ultimately have "f' x = (*v) B"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 68001
diff changeset
  1694
            by (force simp: algebra_simps scalar_mult_eq_scaleR)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1695
          show "matrix (f' x) $ m $ n \<le> b"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1696
          proof (rule tendsto_upperbound [of "\<lambda>i. (A i $ m $ n)" _ sequentially])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1697
            show "(\<lambda>i. A i $ m $ n) \<longlonglongrightarrow> matrix (f' x) $ m $ n"
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
  1698
              by (simp add: B \<open>f' x = (*v) B\<close>)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1699
            show "\<forall>\<^sub>F i in sequentially. A i $ m $ n \<le> b"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1700
              by (simp add: Ab less_eq_real_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1701
          qed auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1702
        next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1703
          fix e :: "real"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1704
          assume "x \<in> S" and b: "matrix (f' x) $ m $ n \<le> b" and "e > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1705
          then obtain d where "d>0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1706
            and d: "\<And>y. y\<in>S \<Longrightarrow> 0 < dist y x \<and> dist y x < d \<longrightarrow> norm (f y - f x - f' x (y - x)) / (norm (y - x))
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1707
                  < e/2"
70802
160eaf566bcb formally augmented corresponding rules for field_simps
haftmann
parents: 70724
diff changeset
  1708
            using f [OF \<open>x \<in> S\<close>]
160eaf566bcb formally augmented corresponding rules for field_simps
haftmann
parents: 70724
diff changeset
  1709
            by (simp add: Deriv.has_derivative_at_within Lim_within)
160eaf566bcb formally augmented corresponding rules for field_simps
haftmann
parents: 70724
diff changeset
  1710
              (auto simp add: field_simps dest: spec [of _ "e/2"])
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1711
          let ?A = "matrix(f' x) - (\<chi> i j. if i = m \<and> j = n then e / 4 else 0)"
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
  1712
          obtain B where BRats: "\<And>i j. B$i$j \<in> \<rat>" and Bo_e6: "onorm((*v) (?A - B)) < e/6"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1713
            using matrix_rational_approximation \<open>e > 0\<close>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1714
            by (metis zero_less_divide_iff zero_less_numeral)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1715
          show "\<exists>d>0. \<exists>A. A $ m $ n < b \<and> (\<forall>i j. A $ i $ j \<in> \<rat>) \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1716
                (\<forall>y\<in>S. norm (y - x) < d \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> e * norm (y - x))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1717
          proof (intro exI conjI ballI allI impI)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1718
            show "d>0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1719
              by (rule \<open>d>0\<close>)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1720
            show "B $ m $ n < b"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1721
            proof -
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
  1722
              have "\<bar>matrix ((*v) (?A - B)) $ m $ n\<bar> \<le> onorm ((*v) (?A - B))"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1723
                using component_le_onorm [OF matrix_vector_mul_linear, of _ m n] by metis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1724
              then show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1725
                using b Bo_e6 by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1726
            qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1727
            show "B $ i $ j \<in> \<rat>" for i j
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1728
              using BRats by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1729
            show "norm (f y - f x - B *v (y - x)) \<le> e * norm (y - x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1730
              if "y \<in> S" and y: "norm (y - x) < d" for y
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1731
            proof (cases "y = x")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1732
              case True then show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1733
                by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1734
            next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1735
              case False
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1736
              have *: "norm(d' - d) \<le> e/2 \<Longrightarrow> norm(y - (x + d')) < e/2 \<Longrightarrow> norm(y - x - d) \<le> e" for d d' e and x y::"real^'n"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1737
                using norm_triangle_le [of "d' - d" "y - (x + d')"] by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1738
              show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1739
              proof (rule *)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1740
                have split246: "\<lbrakk>norm y \<le> e / 6; norm(x - y) \<le> e / 4\<rbrakk> \<Longrightarrow> norm x \<le> e/2" if "e > 0" for e and x y :: "real^'n"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1741
                  using norm_triangle_le [of y "x-y" "e/2"] \<open>e > 0\<close> by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1742
                have "linear (f' x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1743
                  using True f has_derivative_linear by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1744
                then have "norm (f' x (y - x) - B *v (y - x)) = norm ((matrix (f' x) - B) *v (y - x))"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 68001
diff changeset
  1745
                  by (simp add: matrix_vector_mult_diff_rdistrib)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1746
                also have "\<dots> \<le> (e * norm (y - x)) / 2"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1747
                proof (rule split246)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1748
                  have "norm ((?A - B) *v (y - x)) / norm (y - x) \<le> onorm(\<lambda>x. (?A - B) *v x)"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 68001
diff changeset
  1749
                    by (rule le_onorm) auto
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1750
                  also have  "\<dots> < e/6"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1751
                    by (rule Bo_e6)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1752
                  finally have "norm ((?A - B) *v (y - x)) / norm (y - x) < e / 6" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1753
                  then show "norm ((?A - B) *v (y - x)) \<le> e * norm (y - x) / 6"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70802
diff changeset
  1754
                    by (simp add: field_split_simps False)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1755
                  have "norm ((matrix (f' x) - B) *v (y - x) - ((?A - B) *v (y - x))) = norm ((\<chi> i j. if i = m \<and> j = n then e / 4 else 0) *v (y - x))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1756
                    by (simp add: algebra_simps)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1757
                  also have "\<dots> = norm((e/4) *\<^sub>R (y - x)$n *\<^sub>R axis m (1::real))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1758
                  proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1759
                    have "(\<Sum>j\<in>UNIV. (if i = m \<and> j = n then e / 4 else 0) * (y $ j - x $ j)) * 4 = e * (y $ n - x $ n) * axis m 1 $ i" for i
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1760
                    proof (cases "i=m")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1761
                      case True then show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1762
                        by (auto simp: if_distrib [of "\<lambda>z. z * _"] cong: if_cong)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1763
                    next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1764
                      case False then show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1765
                        by (simp add: axis_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1766
                    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1767
                    then have "(\<chi> i j. if i = m \<and> j = n then e / 4 else 0) *v (y - x) = (e/4) *\<^sub>R (y - x)$n *\<^sub>R axis m (1::real)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1768
                      by (auto simp: vec_eq_iff matrix_vector_mult_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1769
                    then show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1770
                      by metis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1771
                  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1772
                  also have "\<dots> \<le> e * norm (y - x) / 4"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  1773
                  proof -
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  1774
                    have "\<bar>y $ n - x $ n\<bar> \<le> norm (y - x)"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  1775
                      by (metis component_le_norm_cart vector_minus_component)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  1776
                    with \<open>e > 0\<close> show ?thesis
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  1777
                      by (simp add: norm_mult abs_mult)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  1778
                  qed
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1779
                  finally show "norm ((matrix (f' x) - B) *v (y - x) - ((?A - B) *v (y - x))) \<le> e * norm (y - x) / 4" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1780
                  show "0 < e * norm (y - x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1781
                    by (simp add: False \<open>e > 0\<close>)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1782
                qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1783
                finally show "norm (f' x (y - x) - B *v (y - x)) \<le> (e * norm (y - x)) / 2" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1784
                show "norm (f y - (f x + f' x (y - x))) < (e * norm (y - x)) / 2"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1785
                  using False d [OF \<open>y \<in> S\<close>] y by (simp add: dist_norm field_simps)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1786
              qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1787
            qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1788
          qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1789
        qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1790
      qed auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1791
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1792
    show "negligible ?M"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1793
      using negligible_subset [OF nN MN] .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1794
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1795
  then show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1796
    by (simp add: borel_measurable_vimage_halfspace_component_le sets_restrict_space_iff assms)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1797
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1798
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1799
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1800
theorem borel_measurable_det_Jacobian:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1801
 fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1802
  assumes S: "S \<in> sets lebesgue" and f: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1803
  shows "(\<lambda>x. det(matrix(f' x))) \<in> borel_measurable (lebesgue_on S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1804
  unfolding det_def
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
  1805
  by (intro measurable) (auto intro: f borel_measurable_partial_derivatives [OF S])
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1806
68001
0a2a1b6507c1 correction of TeX errors and other oversights
paulson <lp15@cam.ac.uk>
parents: 67999
diff changeset
  1807
text\<open>The localisation wrt S uses the same argument for many similar results.\<close>
0a2a1b6507c1 correction of TeX errors and other oversights
paulson <lp15@cam.ac.uk>
parents: 67999
diff changeset
  1808
(*See HOL Light's MEASURABLE_ON_LEBESGUE_MEASURABLE_PREIMAGE_BOREL, etc.*)
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1809
theorem borel_measurable_lebesgue_on_preimage_borel:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1810
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1811
  assumes "S \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1812
  shows "f \<in> borel_measurable (lebesgue_on S) \<longleftrightarrow>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1813
         (\<forall>T. T \<in> sets borel \<longrightarrow> {x \<in> S. f x \<in> T} \<in> sets lebesgue)"
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1814
proof -
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1815
  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"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1816
         if "T \<in> sets borel" for T
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1817
    proof (cases "0 \<in> T")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1818
      case True
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1819
      then have "{x \<in> S. f x \<in> T} = {x. (if x \<in> S then f x else 0) \<in> T} \<inter> S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1820
                "{x. (if x \<in> S then f x else 0) \<in> T} = {x \<in> S. f x \<in> T} \<union> -S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1821
        by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1822
      then show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1823
        by (metis (no_types, lifting) Compl_in_sets_lebesgue assms sets.Int sets.Un)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1824
    next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1825
      case False
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1826
      then have "{x. (if x \<in> S then f x else 0) \<in> T} = {x \<in> S. f x \<in> T}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1827
        by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1828
      then show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1829
        by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1830
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1831
    then show ?thesis
70532
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70380
diff changeset
  1832
      unfolding borel_measurable_lebesgue_preimage_borel borel_measurable_if [OF assms, symmetric]
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1833
      by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1834
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1835
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1836
lemma sets_lebesgue_almost_borel:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1837
  assumes "S \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1838
  obtains B N where "B \<in> sets borel" "negligible N" "B \<union> N = S"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  1839
  by (metis assms negligible_iff_null_sets negligible_subset null_sets_completionI sets_completionE sets_lborel)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1840
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1841
lemma double_lebesgue_sets:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1842
 assumes S: "S \<in> sets lebesgue" and T: "T \<in> sets lebesgue" and fim: "f ` S \<subseteq> T"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1843
 shows "(\<forall>U. U \<in> sets lebesgue \<and> U \<subseteq> T \<longrightarrow> {x \<in> S. f x \<in> U} \<in> sets lebesgue) \<longleftrightarrow>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1844
          f \<in> borel_measurable (lebesgue_on S) \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1845
          (\<forall>U. negligible U \<and> U \<subseteq> T \<longrightarrow> {x \<in> S. f x \<in> U} \<in> sets lebesgue)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1846
         (is "?lhs \<longleftrightarrow> _ \<and> ?rhs")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1847
  unfolding borel_measurable_lebesgue_on_preimage_borel [OF S]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1848
proof (intro iffI allI conjI impI, safe)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1849
  fix V :: "'b set"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1850
  assume *: "\<forall>U. U \<in> sets lebesgue \<and> U \<subseteq> T \<longrightarrow> {x \<in> S. f x \<in> U} \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1851
    and "V \<in> sets borel"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1852
  then have V: "V \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1853
    by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1854
  have "{x \<in> S. f x \<in> V} = {x \<in> S. f x \<in> T \<inter> V}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1855
    using fim by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1856
  also have "{x \<in> S. f x \<in> T \<inter> V} \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1857
    using T V * le_inf_iff by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1858
  finally show "{x \<in> S. f x \<in> V} \<in> sets lebesgue" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1859
next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1860
  fix U :: "'b set"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1861
  assume "\<forall>U. U \<in> sets lebesgue \<and> U \<subseteq> T \<longrightarrow> {x \<in> S. f x \<in> U} \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1862
         "negligible U" "U \<subseteq> T"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1863
  then show "{x \<in> S. f x \<in> U} \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1864
    using negligible_imp_sets by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1865
next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1866
  fix U :: "'b set"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1867
  assume 1 [rule_format]: "(\<forall>T. T \<in> sets borel \<longrightarrow> {x \<in> S. f x \<in> T} \<in> sets lebesgue)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1868
     and 2 [rule_format]: "\<forall>U. negligible U \<and> U \<subseteq> T \<longrightarrow> {x \<in> S. f x \<in> U} \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1869
     and "U \<in> sets lebesgue" "U \<subseteq> T"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1870
  then obtain C N where C: "C \<in> sets borel \<and> negligible N \<and> C \<union> N = U"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  1871
    using sets_lebesgue_almost_borel by metis
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1872
  then have "{x \<in> S. f x \<in> C} \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1873
    by (blast intro: 1)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1874
  moreover have "{x \<in> S. f x \<in> N} \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1875
    using C \<open>U \<subseteq> T\<close> by (blast intro: 2)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1876
  moreover have "{x \<in> S. f x \<in> C \<union> N} = {x \<in> S. f x \<in> C} \<union> {x \<in> S. f x \<in> N}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1877
    by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1878
  ultimately show "{x \<in> S. f x \<in> U} \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1879
    using C by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1880
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1881
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1882
69683
8b3458ca0762 subsection is always %important
immler
parents: 69679
diff changeset
  1883
subsection\<open>Simplest case of Sard's theorem (we don't need continuity of derivative)\<close>
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1884
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1885
lemma Sard_lemma00:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1886
  fixes P :: "'b::euclidean_space set"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1887
  assumes "a \<ge> 0" and a: "a *\<^sub>R i \<noteq> 0" and i: "i \<in> Basis"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1888
    and P: "P \<subseteq> {x. a *\<^sub>R i \<bullet> x = 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1889
    and "0 \<le> m" "0 \<le> e"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1890
 obtains S where "S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1891
            and "{z. norm z \<le> m \<and> (\<exists>t \<in> P. norm(z - t) \<le> e)} \<subseteq> S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1892
            and "measure lebesgue S \<le> (2 * e) * (2 * m) ^ (DIM('b) - 1)"
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1893
proof -
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1894
  have "a > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1895
    using assms by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1896
  let ?v = "(\<Sum>j\<in>Basis. (if j = i then e else m) *\<^sub>R j)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1897
  show thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1898
  proof
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1899
    have "- e \<le> x \<bullet> i" "x \<bullet> i \<le> e"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1900
      if "t \<in> P" "norm (x - t) \<le> e" for x t
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1901
      using \<open>a > 0\<close> that Basis_le_norm [of i "x-t"] P i
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1902
      by (auto simp: inner_commute algebra_simps)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1903
    moreover have "- m \<le> x \<bullet> j" "x \<bullet> j \<le> m"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1904
      if "norm x \<le> m" "t \<in> P" "norm (x - t) \<le> e" "j \<in> Basis" and "j \<noteq> i"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1905
      for x t j
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1906
      using that Basis_le_norm [of j x] by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1907
    ultimately
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1908
    show "{z. norm z \<le> m \<and> (\<exists>t\<in>P. norm (z - t) \<le> e)} \<subseteq> cbox (-?v) ?v"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1909
      by (auto simp: mem_box)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1910
    have *: "\<forall>k\<in>Basis. - ?v \<bullet> k \<le> ?v \<bullet> k"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1911
      using \<open>0 \<le> m\<close> \<open>0 \<le> e\<close> by (auto simp: inner_Basis)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1912
    have 2: "2 ^ DIM('b) = 2 * 2 ^ (DIM('b) - Suc 0)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1913
      by (metis DIM_positive Suc_pred power_Suc)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1914
    show "measure lebesgue (cbox (-?v) ?v) \<le> 2 * e * (2 * m) ^ (DIM('b) - 1)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1915
      using \<open>i \<in> Basis\<close>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1916
      by (simp add: content_cbox [OF *] prod.distrib prod.If_cases Diff_eq [symmetric] 2)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1917
  qed blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1918
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1919
68001
0a2a1b6507c1 correction of TeX errors and other oversights
paulson <lp15@cam.ac.uk>
parents: 67999
diff changeset
  1920
text\<open>As above, but reorienting the vector (HOL Light's @text{GEOM\_BASIS\_MULTIPLE\_TAC})\<close>
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1921
lemma Sard_lemma0:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1922
  fixes P :: "(real^'n::{finite,wellorder}) set"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1923
  assumes "a \<noteq> 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1924
    and P: "P \<subseteq> {x. a \<bullet> x = 0}" and "0 \<le> m" "0 \<le> e"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1925
  obtains S where "S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1926
    and "{z. norm z \<le> m \<and> (\<exists>t \<in> P. norm(z - t) \<le> e)} \<subseteq> S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1927
    and "measure lebesgue S \<le> (2 * e) * (2 * m) ^ (CARD('n) - 1)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1928
proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1929
  obtain T and k::'n where T: "orthogonal_transformation T" and a: "a = T (norm a *\<^sub>R axis k (1::real))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1930
    using rotation_rightward_line by metis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1931
  have Tinv [simp]: "T (inv T x) = x" for x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1932
    by (simp add: T orthogonal_transformation_surj surj_f_inv_f)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1933
  obtain S where S: "S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1934
    and subS: "{z. norm z \<le> m \<and> (\<exists>t \<in> T-`P. norm(z - t) \<le> e)} \<subseteq> S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1935
    and mS: "measure lebesgue S \<le> (2 * e) * (2 * m) ^ (CARD('n) - 1)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1936
  proof (rule Sard_lemma00 [of "norm a" "axis k (1::real)" "T-`P" m e])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1937
    have "norm a *\<^sub>R axis k 1 \<bullet> x = 0" if "T x \<in> P" for x
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  1938
      by (smt (verit, del_insts) P T a mem_Collect_eq orthogonal_transformation_def subset_eq that)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1939
    then show "T -` P \<subseteq> {x. norm a *\<^sub>R axis k 1 \<bullet> x = 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1940
      by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1941
  qed (use assms T in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1942
  show thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1943
  proof
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1944
    show "T ` S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1945
      using S measurable_orthogonal_image T by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1946
    have "{z. norm z \<le> m \<and> (\<exists>t\<in>P. norm (z - t) \<le> e)} \<subseteq> T ` {z. norm z \<le> m \<and> (\<exists>t\<in>T -` P. norm (z - t) \<le> e)}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1947
    proof clarsimp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1948
      fix x t
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  1949
      assume \<section>: "norm x \<le> m" "t \<in> P" "norm (x - t) \<le> e"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1950
      then have "norm (inv T x) \<le> m"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1951
        using orthogonal_transformation_inv [OF T] by (simp add: orthogonal_transformation_norm)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1952
      moreover have "\<exists>t\<in>T -` P. norm (inv T x - t) \<le> e"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  1953
        by (smt (verit, del_insts) T Tinv \<section> linear_diff orthogonal_transformation_def orthogonal_transformation_norm vimage_eq)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1954
      ultimately show "x \<in> T ` {z. norm z \<le> m \<and> (\<exists>t\<in>T -` P. norm (z - t) \<le> e)}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1955
        by force
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1956
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1957
    then show "{z. norm z \<le> m \<and> (\<exists>t\<in>P. norm (z - t) \<le> e)} \<subseteq> T ` S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1958
      using image_mono [OF subS] by (rule order_trans)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1959
    show "measure lebesgue (T ` S) \<le> 2 * e * (2 * m) ^ (CARD('n) - 1)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1960
      using mS T by (simp add: S measure_orthogonal_image)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1961
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1962
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1963
68001
0a2a1b6507c1 correction of TeX errors and other oversights
paulson <lp15@cam.ac.uk>
parents: 67999
diff changeset
  1964
text\<open>As above, but translating the sets (HOL Light's @text{GEN\_GEOM\_ORIGIN\_TAC})\<close>
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1965
lemma Sard_lemma1:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1966
  fixes P :: "(real^'n::{finite,wellorder}) set"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1967
   assumes P: "dim P < CARD('n)" and "0 \<le> m" "0 \<le> e"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1968
 obtains S where "S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1969
            and "{z. norm(z - w) \<le> m \<and> (\<exists>t \<in> P. norm(z - w - t) \<le> e)} \<subseteq> S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1970
            and "measure lebesgue S \<le> (2 * e) * (2 * m) ^ (CARD('n) - 1)"
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1971
proof -
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1972
  obtain a where "a \<noteq> 0" "P \<subseteq> {x. a \<bullet> x = 0}"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 68001
diff changeset
  1973
    using lowdim_subset_hyperplane [of P] P span_base by auto
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1974
  then obtain S where S: "S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1975
    and subS: "{z. norm z \<le> m \<and> (\<exists>t \<in> P. norm(z - t) \<le> e)} \<subseteq> S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1976
    and mS: "measure lebesgue S \<le> (2 * e) * (2 * m) ^ (CARD('n) - 1)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1977
    by (rule Sard_lemma0 [OF _ _ \<open>0 \<le> m\<close> \<open>0 \<le> e\<close>])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1978
  show thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1979
  proof
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1980
    show "(+)w ` S \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1981
      by (metis measurable_translation S)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1982
    show "{z. norm (z - w) \<le> m \<and> (\<exists>t\<in>P. norm (z - w - t) \<le> e)} \<subseteq> (+)w ` S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1983
      using subS by force
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1984
    show "measure lebesgue ((+)w ` S) \<le> 2 * e * (2 * m) ^ (CARD('n) - 1)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1985
      by (metis measure_translation mS)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1986
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1987
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1988
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1989
lemma Sard_lemma2:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1990
  fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n::{finite,wellorder}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1991
  assumes mlen: "CARD('m) \<le> CARD('n)" (is "?m \<le> ?n")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1992
    and "B > 0" "bounded S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1993
    and derS: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1994
    and rank: "\<And>x. x \<in> S \<Longrightarrow> rank(matrix(f' x)) < CARD('n)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1995
    and B: "\<And>x. x \<in> S \<Longrightarrow> onorm(f' x) \<le> B"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1996
  shows "negligible(f ` S)"
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1997
proof -
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1998
  have lin_f': "\<And>x. x \<in> S \<Longrightarrow> linear(f' x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1999
    using derS has_derivative_linear by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2000
  show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2001
  proof (clarsimp simp add: negligible_outer_le)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2002
    fix e :: "real"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2003
    assume "e > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2004
    obtain c where csub: "S \<subseteq> cbox (- (vec c)) (vec c)" and "c > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2005
    proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2006
      obtain b where b: "\<And>x. x \<in> S \<Longrightarrow> norm x \<le> b"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2007
        using \<open>bounded S\<close> by (auto simp: bounded_iff)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2008
      show thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2009
      proof
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2010
        have "- \<bar>b\<bar> - 1 \<le> x $ i \<and> x $ i \<le> \<bar>b\<bar> + 1" if "x \<in> S" for x i
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2011
          using component_le_norm_cart [of x i] b [OF that] by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2012
        then show "S \<subseteq> cbox (- vec (\<bar>b\<bar> + 1)) (vec (\<bar>b\<bar> + 1))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2013
          by (auto simp: mem_box_cart)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2014
      qed auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2015
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2016
    then have box_cc: "box (- (vec c)) (vec c) \<noteq> {}" and cbox_cc: "cbox (- (vec c)) (vec c) \<noteq> {}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2017
      by (auto simp: interval_eq_empty_cart)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2018
    obtain d where "d > 0" "d \<le> B"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2019
             and d: "(d * 2) * (4 * B) ^ (?n - 1) \<le> e / (2*c) ^ ?m / ?m ^ ?m"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2020
      apply (rule that [of "min B (e / (2*c) ^ ?m / ?m ^ ?m / (4 * B) ^ (?n - 1) / 2)"])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2021
      using \<open>B > 0\<close> \<open>c > 0\<close> \<open>e > 0\<close>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2022
      by (simp_all add: divide_simps min_mult_distrib_right)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2023
    have "\<exists>r. 0 < r \<and> r \<le> 1/2 \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2024
              (x \<in> S
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2025
               \<longrightarrow> (\<forall>y. y \<in> S \<and> norm(y - x) < r
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2026
                       \<longrightarrow> norm(f y - f x - f' x (y - x)) \<le> d * norm(y - x)))" for x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2027
    proof (cases "x \<in> S")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2028
      case True
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2029
      then obtain r where "r > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2030
              and "\<And>y. \<lbrakk>y \<in> S; norm (y - x) < r\<rbrakk>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2031
                       \<Longrightarrow> norm (f y - f x - f' x (y - x)) \<le> d * norm (y - x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2032
        using derS \<open>d > 0\<close> by (force simp: has_derivative_within_alt)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2033
      then show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2034
        by (rule_tac x="min r (1/2)" in exI) simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2035
    next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2036
      case False
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2037
      then show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2038
        by (rule_tac x="1/2" in exI) simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2039
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2040
    then obtain r where r12: "\<And>x. 0 < r x \<and> r x \<le> 1/2"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2041
            and r: "\<And>x y. \<lbrakk>x \<in> S; y \<in> S; norm(y - x) < r x\<rbrakk>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2042
                          \<Longrightarrow> norm(f y - f x - f' x (y - x)) \<le> d * norm(y - x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2043
      by metis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2044
    then have ga: "gauge (\<lambda>x. ball x (r x))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2045
      by (auto simp: gauge_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2046
    obtain \<D> where \<D>: "countable \<D>" and sub_cc: "\<Union>\<D> \<subseteq> cbox (- vec c) (vec c)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2047
      and cbox: "\<And>K. K \<in> \<D> \<Longrightarrow> interior K \<noteq> {} \<and> (\<exists>u v. K = cbox u v)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2048
      and djointish: "pairwise (\<lambda>A B. interior A \<inter> interior B = {}) \<D>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2049
      and covered: "\<And>K. K \<in> \<D> \<Longrightarrow> \<exists>x \<in> S \<inter> K. K \<subseteq> ball x (r x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2050
      and close: "\<And>u v. cbox u v \<in> \<D> \<Longrightarrow> \<exists>n. \<forall>i::'m. v $ i - u $ i = 2*c / 2^n"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2051
      and covers: "S \<subseteq> \<Union>\<D>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2052
      apply (rule covering_lemma [OF csub box_cc ga])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2053
      apply (auto simp: Basis_vec_def cart_eq_inner_axis [symmetric])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2054
      done
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2055
    let ?\<mu> = "measure lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2056
    have "\<exists>T. T \<in> lmeasurable \<and> f ` (K \<inter> S) \<subseteq> T \<and> ?\<mu> T \<le> e / (2*c) ^ ?m * ?\<mu> K"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2057
      if "K \<in> \<D>" for K
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2058
    proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2059
      obtain u v where uv: "K = cbox u v"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2060
        using cbox \<open>K \<in> \<D>\<close> by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2061
      then have uv_ne: "cbox u v \<noteq> {}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2062
        using cbox that by fastforce
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2063
      obtain x where x: "x \<in> S \<inter> cbox u v" "cbox u v \<subseteq> ball x (r x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2064
        using \<open>K \<in> \<D>\<close> covered uv by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2065
      then have "dim (range (f' x)) < ?n"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 68001
diff changeset
  2066
        using rank_dim_range [of "matrix (f' x)"] x rank[of x]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 68001
diff changeset
  2067
        by (auto simp: matrix_works scalar_mult_eq_scaleR lin_f')
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2068
      then obtain T where T: "T \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2069
            and subT: "{z. norm(z - f x) \<le> (2 * B) * norm(v - u) \<and> (\<exists>t \<in> range (f' x). norm(z - f x - t) \<le> d * norm(v - u))} \<subseteq> T"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2070
            and measT: "?\<mu> T \<le> (2 * (d * norm(v - u))) * (2 * ((2 * B) * norm(v - u))) ^ (?n - 1)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2071
                        (is "_ \<le> ?DVU")
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2072
        using Sard_lemma1 [of "range (f' x)" "(2 * B) * norm(v - u)" "d * norm(v - u)"]
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2073
        using \<open>B > 0\<close> \<open>d > 0\<close> by auto
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2074
      show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2075
      proof (intro exI conjI)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2076
        have "f ` (K \<inter> S) \<subseteq> {z. norm(z - f x) \<le> (2 * B) * norm(v - u) \<and> (\<exists>t \<in> range (f' x). norm(z - f x - t) \<le> d * norm(v - u))}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2077
          unfolding uv
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2078
        proof (clarsimp simp: mult.assoc, intro conjI)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2079
          fix y
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2080
          assume y: "y \<in> cbox u v" and "y \<in> S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2081
          then have "norm (y - x) < r x"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2082
            by (metis dist_norm mem_ball norm_minus_commute subsetCE x(2))
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2083
          then have le_dyx: "norm (f y - f x - f' x (y - x)) \<le> d * norm (y - x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2084
            using r [of x y] x \<open>y \<in> S\<close> by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2085
          have yx_le: "norm (y - x) \<le> norm (v - u)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2086
          proof (rule norm_le_componentwise_cart)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 68001
diff changeset
  2087
            show "norm ((y - x) $ i) \<le> norm ((v - u) $ i)" for i
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 68001
diff changeset
  2088
              using x y by (force simp: mem_box_cart dest!: spec [where x=i])
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2089
          qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2090
          have *: "\<lbrakk>norm(y - x - z) \<le> d; norm z \<le> B; d \<le> B\<rbrakk> \<Longrightarrow> norm(y - x) \<le> 2 * B"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2091
            for x y z :: "real^'n::_" and d B
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2092
            using norm_triangle_ineq2 [of "y - x" z] by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2093
          show "norm (f y - f x) \<le> 2 * (B * norm (v - u))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2094
          proof (rule * [OF le_dyx])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2095
            have "norm (f' x (y - x)) \<le> onorm (f' x) * norm (y - x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2096
              using onorm [of "f' x" "y-x"] by (meson IntE lin_f' linear_linear x(1))
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2097
            also have "\<dots> \<le> B * norm (v - u)"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2098
              by (meson B IntE lin_f' linear_linear mult_mono' norm_ge_zero onorm_pos_le x(1) yx_le)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2099
            finally show "norm (f' x (y - x)) \<le> B * norm (v - u)" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2100
            show "d * norm (y - x) \<le> B * norm (v - u)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2101
              using \<open>B > 0\<close> by (auto intro: mult_mono [OF \<open>d \<le> B\<close> yx_le])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2102
          qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2103
          show "\<exists>t. norm (f y - f x - f' x t) \<le> d * norm (v - u)"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2104
            by (smt (verit, best) \<open>0 < d\<close> le_dyx mult_le_cancel_left_pos yx_le)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2105
        qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2106
        with subT show "f ` (K \<inter> S) \<subseteq> T" by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2107
        show "?\<mu> T \<le> e / (2*c) ^ ?m * ?\<mu> K"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2108
        proof (rule order_trans [OF measT])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2109
          have "?DVU = (d * 2 * (4 * B) ^ (?n - 1)) * norm (v - u)^?n"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2110
            using \<open>c > 0\<close>
71633
07bec530f02e cleaned proofs
nipkow
parents: 71192
diff changeset
  2111
            apply (simp add: algebra_simps)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2112
            by (metis Suc_pred power_Suc zero_less_card_finite)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2113
          also have "\<dots> \<le> (e / (2*c) ^ ?m / (?m ^ ?m)) * norm(v - u) ^ ?n"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2114
            by (rule mult_right_mono [OF d]) auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2115
          also have "\<dots> \<le> e / (2*c) ^ ?m * ?\<mu> K"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2116
          proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2117
            have "u \<in> ball (x) (r x)" "v \<in> ball x (r x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2118
              using box_ne_empty(1) contra_subsetD [OF x(2)] mem_box(2) uv_ne by fastforce+
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2119
            moreover have "r x \<le> 1/2"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2120
              using r12 by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2121
            ultimately have "norm (v - u) \<le> 1"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2122
              using norm_triangle_half_r [of x u 1 v]
73932
fd21b4a93043 added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents: 73648
diff changeset
  2123
              by (metis (no_types, opaque_lifting) dist_commute dist_norm less_eq_real_def less_le_trans mem_ball)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2124
            then have "norm (v - u) ^ ?n \<le> norm (v - u) ^ ?m"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2125
              by (simp add: power_decreasing [OF mlen])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2126
            also have "\<dots> \<le> ?\<mu> K * real (?m ^ ?m)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2127
            proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2128
              obtain n where n: "\<And>i. v$i - u$i = 2 * c / 2^n"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2129
                using close [of u v] \<open>K \<in> \<D>\<close> uv by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2130
              have "norm (v - u) ^ ?m \<le> (\<Sum>i\<in>UNIV. \<bar>(v - u) $ i\<bar>) ^ ?m"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2131
                by (intro norm_le_l1_cart power_mono) auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2132
              also have "\<dots> \<le> (\<Prod>i\<in>UNIV. v $ i - u $ i) * real CARD('m) ^ CARD('m)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2133
                by (simp add: n field_simps \<open>c > 0\<close> less_eq_real_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2134
              also have "\<dots> = ?\<mu> K * real (?m ^ ?m)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2135
                by (simp add: uv uv_ne content_cbox_cart)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2136
              finally show ?thesis .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2137
            qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2138
            finally have *: "1 / real (?m ^ ?m) * norm (v - u) ^ ?n \<le> ?\<mu> K"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70802
diff changeset
  2139
              by (simp add: field_split_simps)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2140
            show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2141
              using mult_left_mono [OF *, of "e / (2*c) ^ ?m"] \<open>c > 0\<close> \<open>e > 0\<close> by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2142
          qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2143
          finally show "?DVU \<le> e / (2*c) ^ ?m * ?\<mu> K" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2144
        qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2145
      qed (use T in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2146
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2147
    then obtain g where meas_g: "\<And>K. K \<in> \<D> \<Longrightarrow> g K \<in> lmeasurable"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2148
                    and sub_g: "\<And>K. K \<in> \<D> \<Longrightarrow> f ` (K \<inter> S) \<subseteq> g K"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2149
                    and le_g: "\<And>K. K \<in> \<D> \<Longrightarrow> ?\<mu> (g K) \<le> e / (2*c)^?m * ?\<mu> K"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2150
      by metis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2151
    have le_e: "?\<mu> (\<Union>i\<in>\<F>. g i) \<le> e"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2152
      if "\<F> \<subseteq> \<D>" "finite \<F>" for \<F>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2153
    proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2154
      have "?\<mu> (\<Union>i\<in>\<F>. g i) \<le> (\<Sum>i\<in>\<F>. ?\<mu> (g i))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2155
        using meas_g \<open>\<F> \<subseteq> \<D>\<close> by (auto intro: measure_UNION_le [OF \<open>finite \<F>\<close>])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2156
      also have "\<dots> \<le> (\<Sum>K\<in>\<F>. e / (2*c) ^ ?m * ?\<mu> K)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2157
        using \<open>\<F> \<subseteq> \<D>\<close> sum_mono [OF le_g] by (meson le_g subsetCE sum_mono)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2158
      also have "\<dots> = e / (2*c) ^ ?m * (\<Sum>K\<in>\<F>. ?\<mu> K)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2159
        by (simp add: sum_distrib_left)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2160
      also have "\<dots> \<le> e"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2161
      proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2162
        have "\<F> division_of \<Union>\<F>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2163
        proof (rule division_ofI)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2164
          show "K \<subseteq> \<Union>\<F>"  "K \<noteq> {}" "\<exists>a b. K = cbox a b" if "K \<in> \<F>" for K
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2165
            using \<open>K \<in> \<F>\<close> covered cbox \<open>\<F> \<subseteq> \<D>\<close> by (auto simp: Union_upper)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2166
          show "interior K \<inter> interior L = {}" if "K \<in> \<F>" and "L \<in> \<F>" and "K \<noteq> L" for K L
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2167
            by (metis (mono_tags, lifting) \<open>\<F> \<subseteq> \<D>\<close> pairwiseD djointish pairwise_subset that)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2168
        qed (use that in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2169
        then have "sum ?\<mu> \<F> \<le> ?\<mu> (\<Union>\<F>)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2170
          by (simp add: content_division)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2171
        also have "\<dots> \<le> ?\<mu> (cbox (- vec c) (vec c) :: (real, 'm) vec set)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2172
        proof (rule measure_mono_fmeasurable)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2173
          show "\<Union>\<F> \<subseteq> cbox (- vec c) (vec c)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2174
            by (meson Sup_subset_mono sub_cc order_trans \<open>\<F> \<subseteq> \<D>\<close>)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2175
        qed (use \<open>\<F> division_of \<Union>\<F>\<close> lmeasurable_division in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2176
        also have "\<dots> = content (cbox (- vec c) (vec c) :: (real, 'm) vec set)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2177
          by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2178
        also have "\<dots> \<le> (2 ^ ?m * c ^ ?m)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2179
          using \<open>c > 0\<close> by (simp add: content_cbox_if_cart)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2180
        finally have "sum ?\<mu> \<F> \<le> (2 ^ ?m * c ^ ?m)" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2181
        then show ?thesis
71633
07bec530f02e cleaned proofs
nipkow
parents: 71192
diff changeset
  2182
          using \<open>e > 0\<close> \<open>c > 0\<close> by (auto simp: field_split_simps)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2183
      qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2184
      finally show ?thesis .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2185
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2186
    show "\<exists>T. f ` S \<subseteq> T \<and> T \<in> lmeasurable \<and> ?\<mu> T \<le> e"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2187
    proof (intro exI conjI)
69325
4b6ddc5989fc removed legacy input syntax
haftmann
parents: 69313
diff changeset
  2188
      show "f ` S \<subseteq> \<Union> (g ` \<D>)"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2189
        using covers sub_g by force
69325
4b6ddc5989fc removed legacy input syntax
haftmann
parents: 69313
diff changeset
  2190
      show "\<Union> (g ` \<D>) \<in> lmeasurable"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2191
        by (rule fmeasurable_UN_bound [OF \<open>countable \<D>\<close> meas_g le_e])
69325
4b6ddc5989fc removed legacy input syntax
haftmann
parents: 69313
diff changeset
  2192
      show "?\<mu> (\<Union> (g ` \<D>)) \<le> e"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2193
        by (rule measure_UN_bound [OF \<open>countable \<D>\<close> meas_g le_e])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2194
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2195
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2196
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2197
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2198
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  2199
theorem baby_Sard:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2200
  fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n::{finite,wellorder}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2201
  assumes mlen: "CARD('m) \<le> CARD('n)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2202
    and der: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2203
    and rank: "\<And>x. x \<in> S \<Longrightarrow> rank(matrix(f' x)) < CARD('n)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2204
  shows "negligible(f ` S)"
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  2205
proof -
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2206
  let ?U = "\<lambda>n. {x \<in> S. norm(x) \<le> n \<and> onorm(f' x) \<le> real n}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2207
  have "\<And>x. x \<in> S \<Longrightarrow> \<exists>n. norm x \<le> real n \<and> onorm (f' x) \<le> real n"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2208
    by (meson linear order_trans real_arch_simple)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2209
  then have eq: "S = (\<Union>n. ?U n)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2210
    by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2211
  have "negligible (f ` ?U n)" for n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2212
  proof (rule Sard_lemma2 [OF mlen])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2213
    show "0 < real n + 1"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2214
      by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2215
    show "bounded (?U n)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2216
      using bounded_iff by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2217
    show "(f has_derivative f' x) (at x within ?U n)" if "x \<in> ?U n" for x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2218
      using der that by (force intro: has_derivative_subset)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2219
  qed (use rank in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2220
  then show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2221
    by (subst eq) (simp add: image_Union negligible_Union_nat)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2222
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2223
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2224
69683
8b3458ca0762 subsection is always %important
immler
parents: 69679
diff changeset
  2225
subsection\<open>A one-way version of change-of-variables not assuming injectivity. \<close>
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2226
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  2227
lemma integral_on_image_ubound_weak:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2228
  fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2229
  assumes S: "S \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2230
      and f: "f \<in> borel_measurable (lebesgue_on (g ` S))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2231
      and nonneg_fg:  "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f(g x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2232
      and der_g:   "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2233
      and det_int_fg: "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) integrable_on S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2234
      and meas_gim: "\<And>T. \<lbrakk>T \<subseteq> g ` S; T \<in> sets lebesgue\<rbrakk> \<Longrightarrow> {x \<in> S. g x \<in> T} \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2235
    shows "f integrable_on (g ` S) \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2236
           integral (g ` S) f \<le> integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2237
         (is "_ \<and> _ \<le> ?b")
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  2238
proof -
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2239
  let ?D = "\<lambda>x. \<bar>det (matrix (g' x))\<bar>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2240
  have cont_g: "continuous_on S g"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2241
    using der_g has_derivative_continuous_on by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2242
  have [simp]: "space (lebesgue_on S) = S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2243
    by (simp add: S)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2244
  have gS_in_sets_leb: "g ` S \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2245
    apply (rule differentiable_image_in_sets_lebesgue)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2246
    using der_g by (auto simp: S differentiable_def differentiable_on_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2247
  obtain h where nonneg_h: "\<And>n x. 0 \<le> h n x"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2248
    and h_le_f: "\<And>n x. x \<in> S \<Longrightarrow> h n (g x) \<le> f (g x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2249
    and h_inc: "\<And>n x. h n x \<le> h (Suc n) x"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2250
    and h_meas: "\<And>n. h n \<in> borel_measurable lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2251
    and fin_R: "\<And>n. finite(range (h n))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2252
    and lim: "\<And>x. x \<in> g ` S \<Longrightarrow> (\<lambda>n. h n x) \<longlonglongrightarrow> f x"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2253
  proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2254
    let ?f = "\<lambda>x. if x \<in> g ` S then f x else 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2255
    have "?f \<in> borel_measurable lebesgue \<and> (\<forall>x. 0 \<le> ?f x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2256
      by (auto simp: gS_in_sets_leb f nonneg_fg measurable_restrict_space_iff [symmetric])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2257
    then show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2258
      apply (clarsimp simp add: borel_measurable_simple_function_limit_increasing)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2259
      apply (rename_tac h)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2260
      by (rule_tac h=h in that) (auto split: if_split_asm)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2261
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2262
  have h_lmeas: "{t. h n (g t) = y} \<inter> S \<in> sets lebesgue" for y n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2263
  proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2264
    have "space (lebesgue_on (UNIV::(real,'n) vec set)) = UNIV"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2265
      by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2266
    then have "((h n) -`{y} \<inter> g ` S) \<in> sets (lebesgue_on (g ` S))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2267
      by (metis Int_commute borel_measurable_vimage h_meas image_eqI inf_top.right_neutral sets_restrict_space space_borel space_completion space_lborel)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2268
    then have "({u. h n u = y} \<inter> g ` S) \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2269
      using gS_in_sets_leb
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2270
      by (simp add: integral_indicator fmeasurableI2 sets_restrict_space_iff vimage_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2271
    then have "{x \<in> S. g x \<in> ({u. h n u = y} \<inter> g ` S)} \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2272
      using meas_gim[of "({u. h n u = y} \<inter> g ` S)"] by force
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2273
    moreover have "{t. h n (g t) = y} \<inter> S = {x \<in> S. g x \<in> ({u. h n u = y} \<inter> g ` S)}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2274
      by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2275
    ultimately show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2276
      by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2277
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2278
  have hint: "h n integrable_on g ` S \<and> integral (g ` S) (h n) \<le> integral S (\<lambda>x. ?D x * h n (g x))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2279
          (is "?INT \<and> ?lhs \<le> ?rhs")  for n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2280
  proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2281
    let ?R = "range (h n)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2282
    have hn_eq: "h n = (\<lambda>x. \<Sum>y\<in>?R. y * indicat_real {x. h n x = y} x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2283
      by (simp add: indicator_def if_distrib fin_R cong: if_cong)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2284
    have yind: "(\<lambda>t. y * indicator{x. h n x = y} t) integrable_on (g ` S) \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2285
                (integral (g ` S) (\<lambda>t. y * indicator {x. h n x = y} t))
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2286
                 \<le> integral S (\<lambda>t. \<bar>det (matrix (g' t))\<bar> * y * indicator {x. h n x = y} (g t))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2287
       if y: "y \<in> ?R" for y::real
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2288
    proof (cases "y=0")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2289
      case True
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2290
      then show ?thesis using gS_in_sets_leb integrable_0 by force
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2291
    next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2292
      case False
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2293
      with that have "y > 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2294
        using less_eq_real_def nonneg_h by fastforce
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2295
      have "(\<lambda>x. if x \<in> {t. h n (g t) = y} then ?D x else 0) integrable_on S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2296
      proof (rule measurable_bounded_by_integrable_imp_integrable)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2297
        have "(\<lambda>x. ?D x) \<in> borel_measurable (lebesgue_on ({t. h n (g t) = y} \<inter> S))"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2298
        proof -
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2299
          have "(\<lambda>v. det (matrix (g' v))) \<in> borel_measurable (lebesgue_on (S \<inter> {v. h n (g v) = y}))"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2300
            by (metis Int_lower1 S assms(4) borel_measurable_det_Jacobian measurable_restrict_mono)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2301
          then show ?thesis
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2302
            by (simp add: Int_commute)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2303
        qed
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2304
        then have "(\<lambda>x. if x \<in> {t. h n (g t) = y} \<inter> S then ?D x else 0) \<in> borel_measurable lebesgue"
70532
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70380
diff changeset
  2305
          by (rule borel_measurable_if_I [OF _ h_lmeas])
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2306
        then show "(\<lambda>x. if x \<in> {t. h n (g t) = y} then ?D x else 0) \<in> borel_measurable (lebesgue_on S)"
70532
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70380
diff changeset
  2307
          by (simp add: if_if_eq_conj Int_commute borel_measurable_if [OF S, symmetric])
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2308
        show "(\<lambda>x. ?D x *\<^sub>R f (g x) /\<^sub>R y) integrable_on S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2309
          by (rule integrable_cmul) (use det_int_fg in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2310
        show "norm (if x \<in> {t. h n (g t) = y} then ?D x else 0) \<le> ?D x *\<^sub>R f (g x) /\<^sub>R y"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2311
          if "x \<in> S" for x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2312
          using nonneg_h [of n x] \<open>y > 0\<close> nonneg_fg [of x] h_le_f [of x n] that
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70802
diff changeset
  2313
          by (auto simp: divide_simps mult_left_mono)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2314
      qed (use S in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2315
      then have int_det: "(\<lambda>t. \<bar>det (matrix (g' t))\<bar>) integrable_on ({t. h n (g t) = y} \<inter> S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2316
        using integrable_restrict_Int by force
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2317
      have "(g ` ({t. h n (g t) = y} \<inter> S)) \<in> lmeasurable"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2318
        by (blast intro: has_derivative_subset [OF der_g]  measurable_differentiable_image [OF h_lmeas] int_det)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2319
      moreover have "g ` ({t. h n (g t) = y} \<inter> S) = {x. h n x = y} \<inter> g ` S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2320
        by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2321
      moreover have "measure lebesgue (g ` ({t. h n (g t) = y} \<inter> S))
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2322
                     \<le> integral ({t. h n (g t) = y} \<inter> S) (\<lambda>t. \<bar>det (matrix (g' t))\<bar>)"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2323
        by (blast intro: has_derivative_subset [OF der_g] measure_differentiable_image [OF h_lmeas _ int_det])
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2324
      ultimately show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2325
        using \<open>y > 0\<close> integral_restrict_Int [of S "{t. h n (g t) = y}" "\<lambda>t. \<bar>det (matrix (g' t))\<bar> * y"]
71633
07bec530f02e cleaned proofs
nipkow
parents: 71192
diff changeset
  2326
        apply (simp add: integrable_on_indicator integral_indicator)
73536
5131c388a9b0 simplified definition
haftmann
parents: 72569
diff changeset
  2327
        apply (simp add: indicator_def of_bool_def if_distrib cong: if_cong)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2328
        done
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2329
    qed
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2330
    show ?thesis
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2331
    proof
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2332
      show "h n integrable_on g ` S"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2333
        apply (subst hn_eq)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2334
        using yind by (force intro: integrable_sum [OF fin_R])
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2335
      have "?lhs = integral (g ` S) (\<lambda>x. \<Sum>y\<in>range (h n). y * indicat_real {x. h n x = y} x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2336
        by (metis hn_eq)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2337
      also have "\<dots> = (\<Sum>y\<in>range (h n). integral (g ` S) (\<lambda>x. y * indicat_real {x. h n x = y} x))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2338
        by (rule integral_sum [OF fin_R]) (use yind in blast)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2339
      also have "\<dots> \<le> (\<Sum>y\<in>range (h n). integral S (\<lambda>u. \<bar>det (matrix (g' u))\<bar> * y * indicat_real {x. h n x = y} (g u)))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2340
        using yind by (force intro: sum_mono)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2341
      also have "\<dots> = integral S (\<lambda>u. \<Sum>y\<in>range (h n). \<bar>det (matrix (g' u))\<bar> * y * indicat_real {x. h n x = y} (g u))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2342
      proof (rule integral_sum [OF fin_R, symmetric])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2343
        fix y assume y: "y \<in> ?R"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2344
        with nonneg_h have "y \<ge> 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2345
          by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2346
        show "(\<lambda>u. \<bar>det (matrix (g' u))\<bar> * y * indicat_real {x. h n x = y} (g u)) integrable_on S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2347
        proof (rule measurable_bounded_by_integrable_imp_integrable)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2348
          have "(\<lambda>x. indicat_real {x. h n x = y} (g x)) \<in> borel_measurable (lebesgue_on S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2349
            using h_lmeas S
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2350
            by (auto simp: indicator_vimage [symmetric] borel_measurable_indicator_iff sets_restrict_space_iff)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2351
          then show "(\<lambda>u. \<bar>det (matrix (g' u))\<bar> * y * indicat_real {x. h n x = y} (g u)) \<in> borel_measurable (lebesgue_on S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2352
            by (intro borel_measurable_times borel_measurable_abs borel_measurable_const borel_measurable_det_Jacobian [OF S der_g])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2353
        next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2354
          fix x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2355
          assume "x \<in> S"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2356
          then have "y * indicat_real {x. h n x = y} (g x) \<le> f (g x)"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2357
            by (metis (full_types) h_le_f indicator_simps mem_Collect_eq mult.right_neutral mult_zero_right nonneg_fg)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2358
          with \<open>y \<ge> 0\<close> show "norm (?D x * y * indicat_real {x. h n x = y} (g x)) \<le> ?D x * f(g x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2359
            by (simp add: abs_mult mult.assoc mult_left_mono)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2360
        qed (use S det_int_fg in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2361
      qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2362
      also have "\<dots> = integral S (\<lambda>T. \<bar>det (matrix (g' T))\<bar> *
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2363
                                        (\<Sum>y\<in>range (h n). y * indicat_real {x. h n x = y} (g T)))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2364
        by (simp add: sum_distrib_left mult.assoc)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2365
      also have "\<dots> = ?rhs"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2366
        by (metis hn_eq)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2367
      finally show "integral (g ` S) (h n) \<le> ?rhs" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2368
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2369
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2370
  have le: "integral S (\<lambda>T. \<bar>det (matrix (g' T))\<bar> * h n (g T)) \<le> ?b" for n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2371
  proof (rule integral_le)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2372
    show "(\<lambda>T. \<bar>det (matrix (g' T))\<bar> * h n (g T)) integrable_on S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2373
    proof (rule measurable_bounded_by_integrable_imp_integrable)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2374
      have "(\<lambda>T. \<bar>det (matrix (g' T))\<bar> *\<^sub>R h n (g T)) \<in> borel_measurable (lebesgue_on S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2375
      proof (intro borel_measurable_scaleR borel_measurable_abs borel_measurable_det_Jacobian \<open>S \<in> sets lebesgue\<close>)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2376
        have eq: "{x \<in> S. f x \<le> a} = (\<Union>b \<in> (f ` S) \<inter> atMost a. {x. f x = b} \<inter> S)" for f and a::real
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2377
          by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2378
        have "finite ((\<lambda>x. h n (g x)) ` S \<inter> {..a})" for a
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2379
          by (force intro: finite_subset [OF _ fin_R])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2380
        with h_lmeas [of n] show "(\<lambda>x. h n (g x)) \<in> borel_measurable (lebesgue_on S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2381
          apply (simp add: borel_measurable_vimage_halfspace_component_le \<open>S \<in> sets lebesgue\<close> sets_restrict_space_iff eq)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2382
          by (metis (mono_tags) SUP_inf sets.finite_UN)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2383
      qed (use der_g in blast)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2384
      then show "(\<lambda>T. \<bar>det (matrix (g' T))\<bar> * h n (g T)) \<in> borel_measurable (lebesgue_on S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2385
        by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2386
      show "norm (?D x * h n (g x)) \<le> ?D x *\<^sub>R f (g x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2387
        if "x \<in> S" for x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2388
        by (simp add: h_le_f mult_left_mono nonneg_h that)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2389
    qed (use S det_int_fg in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2390
    show "?D x * h n (g x) \<le> ?D x * f (g x)" if "x \<in> S" for x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2391
      by (simp add: \<open>x \<in> S\<close> h_le_f mult_left_mono)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2392
    show "(\<lambda>x. ?D x * f (g x)) integrable_on S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2393
      using det_int_fg by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2394
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2395
  have "f integrable_on g ` S \<and> (\<lambda>k. integral (g ` S) (h k)) \<longlonglongrightarrow> integral (g ` S) f"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2396
  proof (rule monotone_convergence_increasing)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2397
    have "\<bar>integral (g ` S) (h n)\<bar> \<le> integral S (\<lambda>x. ?D x * f (g x))" for n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2398
    proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2399
      have "\<bar>integral (g ` S) (h n)\<bar> = integral (g ` S) (h n)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2400
        using hint by (simp add: integral_nonneg nonneg_h)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2401
      also have "\<dots> \<le> integral S (\<lambda>x. ?D x * f (g x))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2402
        using hint le by (meson order_trans)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2403
      finally show ?thesis .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2404
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2405
    then show "bounded (range (\<lambda>k. integral (g ` S) (h k)))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2406
      by (force simp: bounded_iff)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2407
  qed (use h_inc lim hint in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2408
  moreover have "integral (g ` S) (h n) \<le> integral S (\<lambda>x. ?D x * f (g x))" for n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2409
    using hint by (blast intro: le order_trans)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2410
  ultimately show ?thesis
68532
f8b98d31ad45 Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents: 68403
diff changeset
  2411
    by (auto intro: Lim_bounded)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2412
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2413
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2414
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  2415
lemma integral_on_image_ubound_nonneg:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2416
  fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2417
  assumes nonneg_fg: "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f(g x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2418
      and der_g:   "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2419
      and intS: "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) integrable_on S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2420
  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))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2421
         (is "_ \<and> _ \<le> ?b")
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  2422
proof -
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2423
  let ?D = "\<lambda>x. det (matrix (g' x))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2424
  define S' where "S' \<equiv> {x \<in> S. ?D x * f(g x) \<noteq> 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2425
  then have der_gS': "\<And>x. x \<in> S' \<Longrightarrow> (g has_derivative g' x) (at x within S')"
72445
2c2de074832e tidying and removal of legacy name
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2426
    by (metis (mono_tags, lifting) der_g has_derivative_subset mem_Collect_eq subset_iff)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2427
  have "(\<lambda>x. if x \<in> S then \<bar>?D x\<bar> * f (g x) else 0) integrable_on UNIV"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2428
    by (simp add: integrable_restrict_UNIV intS)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2429
  then have Df_borel: "(\<lambda>x. if x \<in> S then \<bar>?D x\<bar> * f (g x) else 0) \<in> borel_measurable lebesgue"
70707
125705f5965f A little-known material, and some tidying up
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
  2430
    using integrable_imp_measurable lebesgue_on_UNIV_eq by force
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2431
  have S': "S' \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2432
  proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2433
    from Df_borel borel_measurable_vimage_open [of _ UNIV]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2434
    have "{x. (if x \<in> S then \<bar>?D x\<bar> * f (g x) else 0) \<in> T} \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2435
      if "open T" for T
70707
125705f5965f A little-known material, and some tidying up
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
  2436
      using that unfolding lebesgue_on_UNIV_eq
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2437
      by (fastforce simp add: dest!: spec)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2438
    then have "{x. (if x \<in> S then \<bar>?D x\<bar> * f (g x) else 0) \<in> -{0}} \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2439
      using open_Compl by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2440
    then show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2441
      by (simp add: S'_def conj_ac split: if_split_asm cong: conj_cong)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2442
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2443
  then have gS': "g ` S' \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2444
  proof (rule differentiable_image_in_sets_lebesgue)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2445
    show "g differentiable_on S'"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2446
      using der_g unfolding S'_def differentiable_def differentiable_on_def
72445
2c2de074832e tidying and removal of legacy name
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2447
      by (blast intro: has_derivative_subset)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2448
  qed auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2449
  have f: "f \<in> borel_measurable (lebesgue_on (g ` S'))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2450
  proof (clarsimp simp add: borel_measurable_vimage_open)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2451
    fix T :: "real set"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2452
    assume "open T"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2453
    have "{x \<in> g ` S'. f x \<in> T} = g ` {x \<in> S'. f(g x) \<in> T}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2454
      by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2455
    moreover have "g ` {x \<in> S'. f(g x) \<in> T} \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2456
    proof (rule differentiable_image_in_sets_lebesgue)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2457
      let ?h = "\<lambda>x. \<bar>?D x\<bar> * f (g x) /\<^sub>R \<bar>?D x\<bar>"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2458
      have "(\<lambda>x. if x \<in> S' then \<bar>?D x\<bar> * f (g x) else 0) = (\<lambda>x. if x \<in> S then \<bar>?D x\<bar> * f (g x) else 0)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2459
        by (auto simp: S'_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2460
      also have "\<dots> \<in> borel_measurable lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2461
        by (rule Df_borel)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2462
      finally have *: "(\<lambda>x. \<bar>?D x\<bar> * f (g x)) \<in> borel_measurable (lebesgue_on S')"
70532
fcf3b891ccb1 new material; rotated premises of Lim_transform_eventually
paulson <lp15@cam.ac.uk>
parents: 70380
diff changeset
  2463
        by (simp add: borel_measurable_if_D)
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2464
      have "(\<lambda>v. det (matrix (g' v))) \<in> borel_measurable (lebesgue_on S')"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2465
        using S' borel_measurable_det_Jacobian der_gS' by blast
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2466
      then have "?h \<in> borel_measurable (lebesgue_on S')"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2467
        using "*" borel_measurable_abs borel_measurable_inverse borel_measurable_scaleR by blast
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2468
      moreover have "?h x = f(g x)" if "x \<in> S'" for x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2469
        using that by (auto simp: S'_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2470
      ultimately have "(\<lambda>x. f(g x)) \<in> borel_measurable (lebesgue_on S')"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2471
        by (metis (no_types, lifting) measurable_lebesgue_cong)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2472
      then show "{x \<in> S'. f (g x) \<in> T} \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2473
        by (simp add: \<open>S' \<in> sets lebesgue\<close> \<open>open T\<close> borel_measurable_vimage_open sets_restrict_space_iff)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2474
      show "g differentiable_on {x \<in> S'. f (g x) \<in> T}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2475
        using der_g unfolding S'_def differentiable_def differentiable_on_def
72445
2c2de074832e tidying and removal of legacy name
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2476
        by (blast intro: has_derivative_subset)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2477
    qed auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2478
    ultimately have "{x \<in> g ` S'. f x \<in> T} \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2479
      by metis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2480
    then show "{x \<in> g ` S'. f x \<in> T} \<in> sets (lebesgue_on (g ` S'))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2481
      by (simp add: \<open>g ` S' \<in> sets lebesgue\<close> sets_restrict_space_iff)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2482
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2483
  have intS': "(\<lambda>x. \<bar>?D x\<bar> * f (g x)) integrable_on S'"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2484
    using intS
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2485
    by (rule integrable_spike_set) (auto simp: S'_def intro: empty_imp_negligible)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2486
  have lebS': "{x \<in> S'. g x \<in> T} \<in> sets lebesgue" if "T \<subseteq> g ` S'" "T \<in> sets lebesgue" for T
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2487
  proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2488
    have "g \<in> borel_measurable (lebesgue_on S')"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2489
      using der_gS' has_derivative_continuous_on S'
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2490
      by (blast intro: continuous_imp_measurable_on_sets_lebesgue)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2491
    moreover have "{x \<in> S'. g x \<in> U} \<in> sets lebesgue" if "negligible U" "U \<subseteq> g ` S'" for U
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2492
    proof (intro negligible_imp_sets negligible_differentiable_vimage that)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2493
      fix x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2494
      assume x: "x \<in> S'"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2495
      then have "linear (g' x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2496
        using der_gS' has_derivative_linear by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2497
      with x show "inj (g' x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2498
        by (auto simp: S'_def det_nz_iff_inj)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2499
    qed (use der_gS' in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2500
    ultimately show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2501
      using double_lebesgue_sets [OF S' gS' order_refl] that by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2502
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2503
  have int_gS': "f integrable_on g ` S' \<and> integral (g ` S') f \<le> integral S' (\<lambda>x. \<bar>?D x\<bar> * f(g x))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2504
    using integral_on_image_ubound_weak [OF S' f nonneg_fg der_gS' intS' lebS'] S'_def by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2505
  have "negligible (g ` {x \<in> S. det(matrix(g' x)) = 0})"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2506
  proof (rule baby_Sard, simp_all)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2507
    fix x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2508
    assume x: "x \<in> S \<and> det (matrix (g' x)) = 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2509
    then show "(g has_derivative g' x) (at x within {x \<in> S. det (matrix (g' x)) = 0})"
72445
2c2de074832e tidying and removal of legacy name
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2510
      by (metis (no_types, lifting) der_g has_derivative_subset mem_Collect_eq subsetI)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2511
    then show "rank (matrix (g' x)) < CARD('n)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2512
      using det_nz_iff_inj matrix_vector_mul_linear x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2513
      by (fastforce simp add: less_rank_noninjective)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2514
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2515
  then have negg: "negligible (g ` S - g ` {x \<in> S. ?D x \<noteq> 0})"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2516
    by (rule negligible_subset) (auto simp: S'_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2517
  have null: "g ` {x \<in> S. ?D x \<noteq> 0} - g ` S = {}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2518
    by (auto simp: S'_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2519
  let ?F = "{x \<in> S. f (g x) \<noteq> 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2520
  have eq: "g ` S' = g ` ?F \<inter> g ` {x \<in> S. ?D x \<noteq> 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2521
    by (auto simp: S'_def image_iff)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2522
  show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2523
  proof
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2524
    have "((\<lambda>x. if x \<in> g ` ?F then f x else 0) integrable_on g ` {x \<in> S. ?D x \<noteq> 0})"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2525
      using int_gS' eq integrable_restrict_Int [where f=f]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2526
      by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2527
    then have "f integrable_on g ` {x \<in> S. ?D x \<noteq> 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2528
      by (auto simp: image_iff elim!: integrable_eq)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2529
    then show "f integrable_on g ` S"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2530
      using negg null
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2531
      by (auto intro: integrable_spike_set [OF _ empty_imp_negligible negligible_subset])
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2532
    have "integral (g ` S) f = integral (g ` {x \<in> S. ?D x \<noteq> 0}) f"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2533
      using negg by (auto intro: negligible_subset integral_spike_set)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2534
    also have "\<dots> = integral (g ` {x \<in> S. ?D x \<noteq> 0}) (\<lambda>x. if x \<in> g ` ?F then f x else 0)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2535
      by (auto simp: image_iff intro!: integral_cong)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2536
    also have "\<dots> = integral (g ` S') f"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2537
      using  eq integral_restrict_Int by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2538
    also have "\<dots> \<le> integral S' (\<lambda>x. \<bar>?D x\<bar> * f(g x))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2539
      by (metis int_gS')
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2540
    also have "\<dots> \<le> ?b"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2541
      by (rule integral_subset_le [OF _ intS' intS]) (use nonneg_fg S'_def in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2542
    finally show "integral (g ` S) f \<le> ?b" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2543
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2544
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2545
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2546
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  2547
lemma absolutely_integrable_on_image_real:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2548
  fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2549
  assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2550
    and intS: "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) absolutely_integrable_on S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2551
  shows "f absolutely_integrable_on (g ` S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2552
proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2553
  let ?D = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f (g x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2554
  let ?N = "{x \<in> S. f (g x) < 0}" and ?P = "{x \<in> S. f (g x) > 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2555
  have eq: "{x. (if x \<in> S then ?D x else 0) > 0} = {x \<in> S. ?D x > 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2556
           "{x. (if x \<in> S then ?D x else 0) < 0} = {x \<in> S. ?D x < 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2557
    by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2558
  have "?D integrable_on S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2559
    using intS absolutely_integrable_on_def by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2560
  then have "(\<lambda>x. if x \<in> S then ?D x else 0) integrable_on UNIV"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2561
    by (simp add: integrable_restrict_UNIV)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2562
  then have D_borel: "(\<lambda>x. if x \<in> S then ?D x else 0) \<in> borel_measurable (lebesgue_on UNIV)"
70707
125705f5965f A little-known material, and some tidying up
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
  2563
    using integrable_imp_measurable lebesgue_on_UNIV_eq by blast
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2564
  then have Dlt: "{x \<in> S. ?D x < 0} \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2565
    unfolding borel_measurable_vimage_halfspace_component_lt
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2566
    by (drule_tac x=0 in spec) (auto simp: eq)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2567
  from D_borel have Dgt: "{x \<in> S. ?D x > 0} \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2568
    unfolding borel_measurable_vimage_halfspace_component_gt
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2569
    by (drule_tac x=0 in spec) (auto simp: eq)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2570
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2571
  have dfgbm: "?D \<in> borel_measurable (lebesgue_on S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2572
    using intS absolutely_integrable_on_def integrable_imp_measurable by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2573
  have der_gN: "(g has_derivative g' x) (at x within ?N)" if "x \<in> ?N" for x
72445
2c2de074832e tidying and removal of legacy name
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2574
      using der_g has_derivative_subset that by force
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2575
  have "(\<lambda>x. - f x) integrable_on g ` ?N \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2576
         integral (g ` ?N) (\<lambda>x. - f x) \<le> integral ?N (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * - f (g x))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2577
  proof (rule integral_on_image_ubound_nonneg [OF _ der_gN])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2578
    have 1: "?D integrable_on {x \<in> S. ?D x < 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2579
      using Dlt
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2580
      by (auto intro: set_lebesgue_integral_eq_integral [OF set_integrable_subset] intS)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2581
    have "uminus \<circ> (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * - f (g x)) integrable_on ?N"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2582
      by (simp add: o_def mult_less_0_iff empty_imp_negligible integrable_spike_set [OF 1])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2583
    then show "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> * - f (g x)) integrable_on ?N"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2584
      by (simp add: integrable_neg_iff o_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2585
  qed auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2586
  then have "f integrable_on g ` ?N"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2587
    by (simp add: integrable_neg_iff)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2588
  moreover have "g ` ?N = {y \<in> g ` S. f y < 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2589
    by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2590
  ultimately have "f integrable_on {y \<in> g ` S. f y < 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2591
    by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2592
  then have N: "f absolutely_integrable_on {y \<in> g ` S. f y < 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2593
    by (rule absolutely_integrable_absolutely_integrable_ubound) auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2594
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2595
  have der_gP: "(g has_derivative g' x) (at x within ?P)" if "x \<in> ?P" for x
72445
2c2de074832e tidying and removal of legacy name
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2596
      using der_g has_derivative_subset that by force
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2597
    have "f integrable_on g ` ?P \<and> integral (g ` ?P) f \<le> integral ?P ?D"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2598
    proof (rule integral_on_image_ubound_nonneg [OF _ der_gP])
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2599
      show "?D integrable_on ?P"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2600
      proof (rule integrable_spike_set)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2601
        show "?D integrable_on {x \<in> S. 0 < ?D x}"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2602
          using Dgt
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2603
          by (auto intro: set_lebesgue_integral_eq_integral [OF set_integrable_subset] intS)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2604
      qed (auto simp: zero_less_mult_iff empty_imp_negligible)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2605
    qed auto
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2606
  then have "f integrable_on g ` ?P"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2607
    by metis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2608
  moreover have "g ` ?P = {y \<in> g ` S. f y > 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2609
    by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2610
  ultimately have "f integrable_on {y \<in> g ` S. f y > 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2611
    by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2612
  then have P: "f absolutely_integrable_on {y \<in> g ` S. f y > 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2613
    by (rule absolutely_integrable_absolutely_integrable_lbound) auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2614
  have "(\<lambda>x. if x \<in> g ` S \<and> f x < 0 \<or> x \<in> g ` S \<and> 0 < f x then f x else 0) = (\<lambda>x. if x \<in> g ` S then f x else 0)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2615
    by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2616
  then show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2617
    using absolutely_integrable_Un [OF N P] absolutely_integrable_restrict_UNIV [symmetric, where f=f]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2618
    by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2619
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2620
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2621
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  2622
proposition absolutely_integrable_on_image:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2623
  fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2624
  assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2625
    and intS: "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2626
  shows "f absolutely_integrable_on (g ` S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2627
  apply (rule absolutely_integrable_componentwise [OF absolutely_integrable_on_image_real [OF der_g]])
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
  2628
  using absolutely_integrable_component [OF intS]  by auto
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2629
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  2630
proposition integral_on_image_ubound:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2631
  fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2632
  assumes"\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f(g x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2633
    and "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2634
    and "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) integrable_on S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2635
  shows "integral (g ` S) f \<le> integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x))"
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
  2636
  using integral_on_image_ubound_nonneg [OF assms] by simp
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2637
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2638
69683
8b3458ca0762 subsection is always %important
immler
parents: 69679
diff changeset
  2639
subsection\<open>Change-of-variables theorem\<close>
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2640
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2641
text\<open>The classic change-of-variables theorem. We have two versions with quite general hypotheses,
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2642
the first that the transforming function has a continuous inverse, the second that the base set is
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2643
Lebesgue measurable.\<close>
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  2644
lemma cov_invertible_nonneg_le:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2645
  fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2646
  assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2647
    and der_h: "\<And>y. y \<in> T \<Longrightarrow> (h has_derivative h' y) (at y within T)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2648
    and f0: "\<And>y. y \<in> T \<Longrightarrow> 0 \<le> f y"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2649
    and hg: "\<And>x. x \<in> S \<Longrightarrow> g x \<in> T \<and> h(g x) = x"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2650
    and gh: "\<And>y. y \<in> T \<Longrightarrow> h y \<in> S \<and> g(h y) = y"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2651
    and id: "\<And>y. y \<in> T \<Longrightarrow> h' y \<circ> g'(h y) = id"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2652
  shows "f integrable_on T \<and> (integral T f) \<le> b \<longleftrightarrow>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2653
             (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) integrable_on S \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2654
             integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) \<le> b"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2655
        (is "?lhs = ?rhs")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2656
proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2657
  have Teq: "T = g`S" and Seq: "S = h`T"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2658
    using hg gh image_iff by fastforce+
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2659
  have gS: "g differentiable_on S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2660
    by (meson der_g differentiable_def differentiable_on_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2661
  let ?D = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f (g x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2662
  show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2663
  proof
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2664
    assume ?lhs
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2665
    then have fT: "f integrable_on T" and intf: "integral T f \<le> b"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2666
      by blast+
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2667
    show ?rhs
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2668
    proof
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2669
      let ?fgh = "\<lambda>x. \<bar>det (matrix (h' x))\<bar> * (\<bar>det (matrix (g' (h x)))\<bar> * f (g (h x)))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2670
      have ddf: "?fgh x = f x"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2671
        if "x \<in> T" for x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2672
      proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2673
        have "matrix (h' x) ** matrix (g' (h x)) = mat 1"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2674
          by (metis der_g der_h gh has_derivative_linear local.id matrix_compose matrix_id_mat_1 that)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2675
        then have "\<bar>det (matrix (h' x))\<bar> * \<bar>det (matrix (g' (h x)))\<bar> = 1"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2676
          by (metis abs_1 abs_mult det_I det_mul)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2677
        then show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2678
          by (simp add: gh that)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2679
      qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2680
      have "?D integrable_on (h ` T)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2681
      proof (intro set_lebesgue_integral_eq_integral absolutely_integrable_on_image_real)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2682
        show "(\<lambda>x. ?fgh x) absolutely_integrable_on T"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2683
          by (smt (verit, del_insts) abs_absolutely_integrableI_1 ddf f0 fT integrable_eq)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2684
      qed (use der_h in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2685
      with Seq show "(\<lambda>x. ?D x) integrable_on S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2686
        by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2687
      have "integral S (\<lambda>x. ?D x) \<le> integral T (\<lambda>x. ?fgh x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2688
        unfolding Seq
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2689
      proof (rule integral_on_image_ubound)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2690
        show "(\<lambda>x. ?fgh x) integrable_on T"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2691
        using ddf fT integrable_eq by force
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2692
      qed (use f0 gh der_h in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2693
      also have "\<dots> = integral T f"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2694
        by (force simp: ddf intro: integral_cong)
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2695
      finally show "integral S (\<lambda>x. ?D x) \<le> b"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2696
        using intf by linarith 
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2697
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2698
  next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2699
    assume R: ?rhs
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2700
    then have "f integrable_on g ` S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2701
      using der_g f0 hg integral_on_image_ubound_nonneg by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2702
    moreover have "integral (g ` S) f \<le> integral S (\<lambda>x. ?D x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2703
      by (rule integral_on_image_ubound [OF f0 der_g]) (use R Teq in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2704
    ultimately show ?lhs
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2705
      using R by (simp add: Teq)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2706
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2707
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2708
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2709
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  2710
lemma cov_invertible_nonneg_eq:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2711
  fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2712
  assumes "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2713
      and "\<And>y. y \<in> T \<Longrightarrow> (h has_derivative h' y) (at y within T)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2714
      and "\<And>y. y \<in> T \<Longrightarrow> 0 \<le> f y"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2715
      and "\<And>x. x \<in> S \<Longrightarrow> g x \<in> T \<and> h(g x) = x"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2716
      and "\<And>y. y \<in> T \<Longrightarrow> h y \<in> S \<and> g(h y) = y"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2717
      and "\<And>y. y \<in> T \<Longrightarrow> h' y \<circ> g'(h y) = id"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2718
  shows "((\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) has_integral b) S \<longleftrightarrow> (f has_integral b) T"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2719
  using cov_invertible_nonneg_le [OF assms]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2720
  by (simp add: has_integral_iff) (meson eq_iff)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2721
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2722
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  2723
lemma cov_invertible_real:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2724
  fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real" and g :: "real^'n::_ \<Rightarrow> real^'n::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2725
  assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2726
      and der_h: "\<And>y. y \<in> T \<Longrightarrow> (h has_derivative h' y) (at y within T)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2727
      and hg: "\<And>x. x \<in> S \<Longrightarrow> g x \<in> T \<and> h(g x) = x"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2728
      and gh: "\<And>y. y \<in> T \<Longrightarrow> h y \<in> S \<and> g(h y) = y"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2729
      and id: "\<And>y. y \<in> T \<Longrightarrow> h' y \<circ> g'(h y) = id"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2730
  shows "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) absolutely_integrable_on S \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2731
           integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)) = b \<longleftrightarrow>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2732
         f absolutely_integrable_on T \<and> integral T f = b"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2733
         (is "?lhs = ?rhs")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2734
proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2735
  have Teq: "T = g`S" and Seq: "S = h`T"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2736
    using hg gh image_iff by fastforce+
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2737
  let ?DP = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x)" and ?DN = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> * -f(g x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2738
  have "+": "(?DP has_integral b) {x \<in> S. f (g x) > 0} \<longleftrightarrow> (f has_integral b) {y \<in> T. f y > 0}" for b
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2739
  proof (rule cov_invertible_nonneg_eq)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2740
    have *: "(\<lambda>x. f (g x)) -` Y \<inter> {x \<in> S. f (g x) > 0}
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2741
          = ((\<lambda>x. f (g x)) -` Y \<inter> S) \<inter> {x \<in> S. f (g x) > 0}" for Y
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2742
      by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2743
    show "(g has_derivative g' x) (at x within {x \<in> S. f (g x) > 0})" if "x \<in> {x \<in> S. f (g x) > 0}" for x
72445
2c2de074832e tidying and removal of legacy name
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2744
      using that der_g has_derivative_subset by fastforce
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2745
    show "(h has_derivative h' y) (at y within {y \<in> T. f y > 0})" if "y \<in> {y \<in> T. f y > 0}" for y
72445
2c2de074832e tidying and removal of legacy name
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2746
      using that der_h has_derivative_subset by fastforce
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2747
  qed (use gh hg id in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2748
  have "-": "(?DN has_integral b) {x \<in> S. f (g x) < 0} \<longleftrightarrow> ((\<lambda>x. - f x) has_integral b) {y \<in> T. f y < 0}" for b
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2749
  proof (rule cov_invertible_nonneg_eq)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2750
    have *: "(\<lambda>x. - f (g x)) -` y \<inter> {x \<in> S. f (g x) < 0}
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2751
          = ((\<lambda>x. f (g x)) -` uminus ` y \<inter> S) \<inter> {x \<in> S. f (g x) < 0}" for y
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2752
      using image_iff by fastforce
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2753
    show "(g has_derivative g' x) (at x within {x \<in> S. f (g x) < 0})" if "x \<in> {x \<in> S. f (g x) < 0}" for x
72445
2c2de074832e tidying and removal of legacy name
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2754
      using that der_g has_derivative_subset by fastforce
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2755
    show "(h has_derivative h' y) (at y within {y \<in> T. f y < 0})" if "y \<in> {y \<in> T. f y < 0}" for y
72445
2c2de074832e tidying and removal of legacy name
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2756
      using that der_h has_derivative_subset by fastforce
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2757
  qed (use gh hg id in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2758
  show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2759
  proof
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2760
    assume LHS: ?lhs
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2761
    have eq: "{x. (if x \<in> S then ?DP x else 0) > 0} = {x \<in> S. ?DP x > 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2762
      "{x. (if x \<in> S then ?DP x else 0) < 0} = {x \<in> S. ?DP x < 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2763
      by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2764
    have "?DP integrable_on S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2765
      using LHS absolutely_integrable_on_def by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2766
    then have "(\<lambda>x. if x \<in> S then ?DP x else 0) integrable_on UNIV"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2767
      by (simp add: integrable_restrict_UNIV)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2768
    then have D_borel: "(\<lambda>x. if x \<in> S then ?DP x else 0) \<in> borel_measurable (lebesgue_on UNIV)"
70707
125705f5965f A little-known material, and some tidying up
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
  2769
      using integrable_imp_measurable lebesgue_on_UNIV_eq by blast
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2770
    then have SN: "{x \<in> S. ?DP x < 0} \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2771
      unfolding borel_measurable_vimage_halfspace_component_lt
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2772
      by (drule_tac x=0 in spec) (auto simp: eq)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2773
    from D_borel have SP: "{x \<in> S. ?DP x > 0} \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2774
      unfolding borel_measurable_vimage_halfspace_component_gt
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2775
      by (drule_tac x=0 in spec) (auto simp: eq)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2776
    have "?DP absolutely_integrable_on {x \<in> S. ?DP x > 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2777
      using LHS by (fast intro!: set_integrable_subset [OF _, of _ S] SP)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2778
    then have aP: "?DP absolutely_integrable_on {x \<in> S. f (g x) > 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2779
      by (rule absolutely_integrable_spike_set) (auto simp: zero_less_mult_iff empty_imp_negligible)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2780
    have "?DP absolutely_integrable_on {x \<in> S. ?DP x < 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2781
      using LHS by (fast intro!: set_integrable_subset [OF _, of _ S] SN)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2782
    then have aN: "?DP absolutely_integrable_on {x \<in> S. f (g x) < 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2783
      by (rule absolutely_integrable_spike_set) (auto simp: mult_less_0_iff empty_imp_negligible)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2784
    have fN: "f integrable_on {y \<in> T. f y < 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2785
             "integral {y \<in> T. f y < 0} f = integral {x \<in> S. f (g x) < 0} ?DP"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2786
      using "-" [of "integral {x \<in> S. f(g x) < 0} ?DN"] aN
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2787
      by (auto simp: set_lebesgue_integral_eq_integral has_integral_iff integrable_neg_iff)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2788
    have faN: "f absolutely_integrable_on {y \<in> T. f y < 0}"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2789
    proof (rule absolutely_integrable_integrable_bound)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2790
      show "(\<lambda>x. - f x) integrable_on {y \<in> T. f y < 0}"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2791
        using fN by (auto simp: integrable_neg_iff)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2792
    qed (use fN in auto)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2793
    have fP: "f integrable_on {y \<in> T. f y > 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2794
             "integral {y \<in> T. f y > 0} f = integral {x \<in> S. f (g x) > 0} ?DP"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2795
      using "+" [of "integral {x \<in> S. f(g x) > 0} ?DP"] aP
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2796
      by (auto simp: set_lebesgue_integral_eq_integral has_integral_iff integrable_neg_iff)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2797
    have faP: "f absolutely_integrable_on {y \<in> T. f y > 0}"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2798
      using fP(1) nonnegative_absolutely_integrable_1 by fastforce
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2799
    have fa: "f absolutely_integrable_on ({y \<in> T. f y < 0} \<union> {y \<in> T. f y > 0})"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2800
      by (rule absolutely_integrable_Un [OF faN faP])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2801
    show ?rhs
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2802
    proof
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2803
      have eq: "((if x \<in> T \<and> f x < 0 \<or> x \<in> T \<and> 0 < f x then 1 else 0) * f x)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2804
              = (if x \<in> T then 1 else 0) * f x" for x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2805
        by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2806
      show "f absolutely_integrable_on T"
73536
5131c388a9b0 simplified definition
haftmann
parents: 72569
diff changeset
  2807
        using fa by (simp add: indicator_def of_bool_def set_integrable_def eq)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2808
      have [simp]: "{y \<in> T. f y < 0} \<inter> {y \<in> T. 0 < f y} = {}" for T and f :: "(real^'n::_) \<Rightarrow> real"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2809
        by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2810
      have "integral T f = integral ({y \<in> T. f y < 0} \<union> {y \<in> T. f y > 0}) f"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2811
        by (intro empty_imp_negligible integral_spike_set) (auto simp: eq)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2812
      also have "\<dots> = integral {y \<in> T. f y < 0} f + integral {y \<in> T. f y > 0} f"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2813
        using fN fP by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2814
      also have "\<dots> = integral {x \<in> S. f (g x) < 0} ?DP + integral {x \<in> S. 0 < f (g x)} ?DP"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2815
        by (simp add: fN fP)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2816
      also have "\<dots> = integral ({x \<in> S. f (g x) < 0} \<union> {x \<in> S. 0 < f (g x)}) ?DP"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2817
        using aP aN by (simp add: set_lebesgue_integral_eq_integral)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2818
      also have "\<dots> = integral S ?DP"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2819
        by (intro empty_imp_negligible integral_spike_set) auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2820
      also have "\<dots> = b"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2821
        using LHS by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2822
      finally show "integral T f = b" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2823
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2824
  next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2825
    assume RHS: ?rhs
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2826
    have eq: "{x. (if x \<in> T then f x else 0) > 0} = {x \<in> T. f x > 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2827
             "{x. (if x \<in> T then f x else 0) < 0} = {x \<in> T. f x < 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2828
      by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2829
    have "f integrable_on T"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2830
      using RHS absolutely_integrable_on_def by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2831
    then have "(\<lambda>x. if x \<in> T then f x else 0) integrable_on UNIV"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2832
      by (simp add: integrable_restrict_UNIV)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2833
    then have D_borel: "(\<lambda>x. if x \<in> T then f x else 0) \<in> borel_measurable (lebesgue_on UNIV)"
70707
125705f5965f A little-known material, and some tidying up
paulson <lp15@cam.ac.uk>
parents: 70547
diff changeset
  2834
      using integrable_imp_measurable lebesgue_on_UNIV_eq by blast
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2835
    then have TN: "{x \<in> T. f x < 0} \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2836
      unfolding borel_measurable_vimage_halfspace_component_lt
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2837
      by (drule_tac x=0 in spec) (auto simp: eq)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2838
    from D_borel have TP: "{x \<in> T. f x > 0} \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2839
      unfolding borel_measurable_vimage_halfspace_component_gt
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2840
      by (drule_tac x=0 in spec) (auto simp: eq)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2841
    have aint: "f absolutely_integrable_on {y. y \<in> T \<and> 0 < (f y)}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2842
               "f absolutely_integrable_on {y. y \<in> T \<and> (f y) < 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2843
         and intT: "integral T f = b"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  2844
      using set_integrable_subset [of _ T] TP TN RHS by blast+
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2845
    show ?lhs
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2846
    proof
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2847
      have fN: "f integrable_on {v \<in> T. f v < 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2848
        using absolutely_integrable_on_def aint by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2849
      then have DN: "(?DN has_integral integral {y \<in> T. f y < 0} (\<lambda>x. - f x)) {x \<in> S. f (g x) < 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2850
        using "-" [of "integral {y \<in> T. f y < 0} (\<lambda>x. - f x)"]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2851
        by (simp add: has_integral_neg_iff integrable_integral)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2852
      have aDN: "?DP absolutely_integrable_on {x \<in> S. f (g x) < 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2853
        apply (rule absolutely_integrable_integrable_bound [where g = ?DN])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2854
        using DN hg by (fastforce simp: abs_mult integrable_neg_iff)+
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2855
      have fP: "f integrable_on {v \<in> T. f v > 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2856
        using absolutely_integrable_on_def aint by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2857
      then have DP: "(?DP has_integral integral {y \<in> T. f y > 0} f) {x \<in> S. f (g x) > 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2858
        using "+" [of "integral {y \<in> T. f y > 0} f"]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2859
        by (simp add: has_integral_neg_iff integrable_integral)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2860
      have aDP: "?DP absolutely_integrable_on {x \<in> S. f (g x) > 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2861
        apply (rule absolutely_integrable_integrable_bound [where g = ?DP])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2862
        using DP hg by (fastforce simp: integrable_neg_iff)+
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2863
      have eq: "(if x \<in> S then 1 else 0) * ?DP x = (if x \<in> S \<and> f (g x) < 0 \<or> x \<in> S \<and> f (g x) > 0 then 1 else 0) * ?DP x" for x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2864
        by force
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2865
      have "?DP absolutely_integrable_on ({x \<in> S. f (g x) < 0} \<union> {x \<in> S. f (g x) > 0})"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2866
        by (rule absolutely_integrable_Un [OF aDN aDP])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2867
      then show I: "?DP absolutely_integrable_on S"
73536
5131c388a9b0 simplified definition
haftmann
parents: 72569
diff changeset
  2868
        by (simp add: indicator_def of_bool_def eq set_integrable_def)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2869
      have [simp]: "{y \<in> S. f y < 0} \<inter> {y \<in> S. 0 < f y} = {}" for S and f :: "(real^'n::_) \<Rightarrow> real"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2870
        by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2871
      have "integral S ?DP = integral ({x \<in> S. f (g x) < 0} \<union> {x \<in> S. f (g x) > 0}) ?DP"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2872
        by (intro empty_imp_negligible integral_spike_set) auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2873
      also have "\<dots> = integral {x \<in> S. f (g x) < 0} ?DP + integral {x \<in> S. 0 < f (g x)} ?DP"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2874
        using aDN aDP by (simp add: set_lebesgue_integral_eq_integral)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2875
      also have "\<dots> = - integral {y \<in> T. f y < 0} (\<lambda>x. - f x) + integral {y \<in> T. f y > 0} f"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2876
        using DN DP by (auto simp: has_integral_iff)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2877
      also have "\<dots> = integral ({x \<in> T. f x < 0} \<union> {x \<in> T. 0 < f x}) f"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2878
        by (simp add: fN fP)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2879
      also have "\<dots> = integral T f"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2880
        by (intro empty_imp_negligible integral_spike_set) auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2881
      also have "\<dots> = b"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2882
        using intT by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2883
      finally show "integral S ?DP = b" .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2884
    qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2885
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2886
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2887
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2888
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  2889
lemma cv_inv_version3:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2890
  fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2891
  assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2892
    and der_h: "\<And>y. y \<in> T \<Longrightarrow> (h has_derivative h' y) (at y within T)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2893
    and hg: "\<And>x. x \<in> S \<Longrightarrow> g x \<in> T \<and> h(g x) = x"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2894
    and gh: "\<And>y. y \<in> T \<Longrightarrow> h y \<in> S \<and> g(h y) = y"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2895
    and id: "\<And>y. y \<in> T \<Longrightarrow> h' y \<circ> g'(h y) = id"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2896
  shows "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2897
             integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) = b
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2898
         \<longleftrightarrow> f absolutely_integrable_on T \<and> integral T f = b"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2899
proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2900
  let ?D = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2901
  have "((\<lambda>x. \<bar>det (matrix (g' x))\<bar> * f(g x) $ i) absolutely_integrable_on S \<and> integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> * (f(g x) $ i)) = b $ i) \<longleftrightarrow>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2902
        ((\<lambda>x. f x $ i) absolutely_integrable_on T \<and> integral T (\<lambda>x. f x $ i) = b $ i)" for i
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2903
    by (rule cov_invertible_real [OF der_g der_h hg gh id])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2904
  then have "?D absolutely_integrable_on S \<and> (?D has_integral b) S \<longleftrightarrow>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2905
        f absolutely_integrable_on T \<and> (f has_integral b) T"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2906
    unfolding absolutely_integrable_componentwise_iff [where f=f] has_integral_componentwise_iff [of f]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2907
              absolutely_integrable_componentwise_iff [where f="?D"] has_integral_componentwise_iff [of ?D]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2908
    by (auto simp: all_conj_distrib Basis_vec_def cart_eq_inner_axis [symmetric]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2909
           has_integral_iff set_lebesgue_integral_eq_integral)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2910
  then show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2911
    using absolutely_integrable_on_def by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2912
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2913
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2914
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  2915
lemma cv_inv_version4:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2916
  fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2917
  assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S) \<and> invertible(matrix(g' x))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2918
    and hg: "\<And>x. x \<in> S \<Longrightarrow> continuous_on (g ` S) h \<and> h(g x) = x"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2919
  shows "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2920
             integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) = b
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2921
         \<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2922
proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2923
  have "\<forall>x. \<exists>h'. x \<in> S
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2924
           \<longrightarrow> (g has_derivative g' x) (at x within S) \<and> linear h' \<and> g' x \<circ> h' = id \<and> h' \<circ> g' x = id"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2925
    using der_g matrix_invertible has_derivative_linear by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2926
  then obtain h' where h':
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2927
    "\<And>x. x \<in> S
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2928
           \<Longrightarrow> (g has_derivative g' x) (at x within S) \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2929
               linear (h' x) \<and> g' x \<circ> (h' x) = id \<and> (h' x) \<circ> g' x = id"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2930
    by metis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2931
  show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2932
  proof (rule cv_inv_version3)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2933
    show "\<And>y. y \<in> g ` S \<Longrightarrow> (h has_derivative h' (h y)) (at y within g ` S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2934
      using h' hg
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2935
      by (force simp: continuous_on_eq_continuous_within intro!: has_derivative_inverse_within)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2936
  qed (use h' hg in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2937
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2938
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2939
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  2940
theorem has_absolute_integral_change_of_variables_invertible:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2941
  fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2942
  assumes der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2943
      and hg: "\<And>x. x \<in> S \<Longrightarrow> h(g x) = x"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2944
      and conth: "continuous_on (g ` S) h"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2945
  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>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2946
         f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2947
    (is "?lhs = ?rhs")
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  2948
proof -
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2949
  let ?S = "{x \<in> S. invertible (matrix (g' x))}" and ?D = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2950
  have *: "?D absolutely_integrable_on ?S \<and> integral ?S ?D = b
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2951
           \<longleftrightarrow> f absolutely_integrable_on (g ` ?S) \<and> integral (g ` ?S) f = b"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2952
  proof (rule cv_inv_version4)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2953
    show "(g has_derivative g' x) (at x within ?S) \<and> invertible (matrix (g' x))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2954
      if "x \<in> ?S" for x
72445
2c2de074832e tidying and removal of legacy name
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2955
      using der_g that has_derivative_subset that by fastforce
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2956
    show "continuous_on (g ` ?S) h \<and> h (g x) = x"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2957
      if "x \<in> ?S" for x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2958
      using that continuous_on_subset [OF conth]  by (simp add: hg image_mono)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2959
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2960
  have "(g has_derivative g' x) (at x within {x \<in> S. rank (matrix (g' x)) < CARD('m)})" if "x \<in> S" for x
72445
2c2de074832e tidying and removal of legacy name
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2961
    by (metis (no_types, lifting) der_g has_derivative_subset mem_Collect_eq subsetI that)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2962
  then have "negligible (g ` {x \<in> S. \<not> invertible (matrix (g' x))})"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2963
    by (auto simp: invertible_det_nz det_eq_0_rank intro: baby_Sard)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2964
  then have neg: "negligible {x \<in> g ` S. x \<notin> g ` ?S \<and> f x \<noteq> 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2965
    by (auto intro: negligible_subset)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2966
  have [simp]: "{x \<in> g ` ?S. x \<notin> g ` S \<and> f x \<noteq> 0} = {}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2967
    by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2968
  have "?D absolutely_integrable_on ?S \<and> integral ?S ?D = b
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2969
    \<longleftrightarrow> ?D absolutely_integrable_on S \<and> integral S ?D = b"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2970
    apply (intro conj_cong absolutely_integrable_spike_set_eq)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2971
      apply(auto simp: integral_spike_set invertible_det_nz empty_imp_negligible neg)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2972
    done
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2973
  moreover
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2974
  have "f absolutely_integrable_on (g ` ?S) \<and> integral (g ` ?S) f = b
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2975
    \<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2976
    by (auto intro!: conj_cong absolutely_integrable_spike_set_eq integral_spike_set neg)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2977
  ultimately
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2978
  show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2979
    using "*" by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2980
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2981
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2982
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2983
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  2984
theorem has_absolute_integral_change_of_variables_compact:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2985
  fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2986
  assumes "compact S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2987
      and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2988
      and inj: "inj_on g S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2989
  shows "((\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2990
             integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) = b
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2991
      \<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2992
proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2993
  obtain h where hg: "\<And>x. x \<in> S \<Longrightarrow> h(g x) = x"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2994
    using inj by (metis the_inv_into_f_f)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2995
  have conth: "continuous_on (g ` S) h"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2996
    by (metis \<open>compact S\<close> continuous_on_inv der_g has_derivative_continuous_on hg)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2997
  show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2998
    by (rule has_absolute_integral_change_of_variables_invertible [OF der_g hg conth])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2999
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3000
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3001
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  3002
lemma has_absolute_integral_change_of_variables_compact_family:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3003
  fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3004
  assumes compact: "\<And>n::nat. compact (F n)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3005
      and der_g: "\<And>x. x \<in> (\<Union>n. F n) \<Longrightarrow> (g has_derivative g' x) (at x within (\<Union>n. F n))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3006
      and inj: "inj_on g (\<Union>n. F n)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3007
  shows "((\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on (\<Union>n. F n) \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3008
             integral (\<Union>n. F n) (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) = b
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3009
      \<longleftrightarrow> f absolutely_integrable_on (g ` (\<Union>n. F n)) \<and> integral (g ` (\<Union>n. F n)) f = b)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3010
proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3011
  let ?D = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f (g x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3012
  let ?U = "\<lambda>n. \<Union>m\<le>n. F m"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3013
  let ?lift = "vec::real\<Rightarrow>real^1"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3014
  have F_leb: "F m \<in> sets lebesgue" for m
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3015
    by (simp add: compact borel_compact)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3016
  have iff: "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f (g x)) absolutely_integrable_on (?U n) \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3017
             integral (?U n) (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f (g x)) = b
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3018
         \<longleftrightarrow> f absolutely_integrable_on (g ` (?U n)) \<and> integral (g ` (?U n)) f = b" for n b and f :: "real^'m::_ \<Rightarrow> real^'k"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3019
  proof (rule has_absolute_integral_change_of_variables_compact)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3020
    show "compact (?U n)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3021
      by (simp add: compact compact_UN)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3022
    show "(g has_derivative g' x) (at x within (?U n))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3023
      if "x \<in> ?U n" for x
72445
2c2de074832e tidying and removal of legacy name
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3024
      using that by (blast intro!: has_derivative_subset [OF der_g])
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3025
    show "inj_on g (?U n)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3026
      using inj by (auto simp: inj_on_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3027
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3028
  show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3029
    unfolding image_UN
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3030
  proof safe
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3031
    assume DS: "?D absolutely_integrable_on (\<Union>n. F n)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3032
      and b: "b = integral (\<Union>n. F n) ?D"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3033
    have DU: "\<And>n. ?D absolutely_integrable_on (?U n)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3034
             "(\<lambda>n. integral (?U n) ?D) \<longlonglongrightarrow> integral (\<Union>n. F n) ?D"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3035
      using integral_countable_UN [OF DS F_leb] by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3036
    with iff have fag: "f absolutely_integrable_on g ` (?U n)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3037
      and fg_int: "integral (\<Union>m\<le>n. g ` F m) f = integral (?U n) ?D" for n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3038
      by (auto simp: image_UN)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3039
    let ?h = "\<lambda>x. if x \<in> (\<Union>m. g ` F m) then norm(f x) else 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3040
    have "(\<lambda>x. if x \<in> (\<Union>m. g ` F m) then f x else 0) absolutely_integrable_on UNIV"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3041
    proof (rule dominated_convergence_absolutely_integrable)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3042
      show "(\<lambda>x. if x \<in> (\<Union>m\<le>k. g ` F m) then f x else 0) absolutely_integrable_on UNIV" for k
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3043
        unfolding absolutely_integrable_restrict_UNIV
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3044
        using fag by (simp add: image_UN)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3045
      let ?nf = "\<lambda>n x. if x \<in> (\<Union>m\<le>n. g ` F m) then norm(f x) else 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3046
      show "?h integrable_on UNIV"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3047
      proof (rule monotone_convergence_increasing [THEN conjunct1])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3048
        show "?nf k integrable_on UNIV" for k
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3049
          using fag
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3050
          unfolding integrable_restrict_UNIV absolutely_integrable_on_def by (simp add: image_UN)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3051
        { fix n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3052
          have "(norm \<circ> ?D) absolutely_integrable_on ?U n"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3053
            by (intro absolutely_integrable_norm DU)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3054
          then have "integral (g ` ?U n) (norm \<circ> f) = integral (?U n) (norm \<circ> ?D)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3055
            using iff [of n "vec \<circ> norm \<circ> f" "integral (?U n) (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R (?lift \<circ> norm \<circ> f) (g x))"]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3056
            unfolding absolutely_integrable_on_1_iff integral_on_1_eq by (auto simp: o_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3057
        }
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3058
        moreover have "bounded (range (\<lambda>k. integral (?U k) (norm \<circ> ?D)))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3059
          unfolding bounded_iff
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3060
        proof (rule exI, clarify)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3061
          fix k
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3062
          show "norm (integral (?U k) (norm \<circ> ?D)) \<le> integral (\<Union>n. F n) (norm \<circ> ?D)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3063
            unfolding integral_restrict_UNIV [of _ "norm \<circ> ?D", symmetric]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3064
          proof (rule integral_norm_bound_integral)
69325
4b6ddc5989fc removed legacy input syntax
haftmann
parents: 69313
diff changeset
  3065
            show "(\<lambda>x. if x \<in> \<Union> (F ` {..k}) then (norm \<circ> ?D) x else 0) integrable_on UNIV"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3066
              "(\<lambda>x. if x \<in> (\<Union>n. F n) then (norm \<circ> ?D) x else 0) integrable_on UNIV"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3067
              using DU(1) DS
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3068
              unfolding absolutely_integrable_on_def o_def integrable_restrict_UNIV by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3069
          qed auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3070
        qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3071
        ultimately show "bounded (range (\<lambda>k. integral UNIV (?nf k)))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3072
          by (simp add: integral_restrict_UNIV image_UN [symmetric] o_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3073
      next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3074
        show "(\<lambda>k. if x \<in> (\<Union>m\<le>k. g ` F m) then norm (f x) else 0)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3075
              \<longlonglongrightarrow> (if x \<in> (\<Union>m. g ` F m) then norm (f x) else 0)" for x
70365
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70139
diff changeset
  3076
          by (force intro: tendsto_eventually eventually_sequentiallyI)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3077
      qed auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3078
    next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3079
      show "(\<lambda>k. if x \<in> (\<Union>m\<le>k. g ` F m) then f x else 0)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3080
            \<longlonglongrightarrow> (if x \<in> (\<Union>m. g ` F m) then f x else 0)" for x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3081
      proof clarsimp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3082
        fix m y
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3083
        assume "y \<in> F m"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3084
        show "(\<lambda>k. if \<exists>x\<in>{..k}. g y \<in> g ` F x then f (g y) else 0) \<longlonglongrightarrow> f (g y)"
70365
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70139
diff changeset
  3085
          using \<open>y \<in> F m\<close> by (force intro: tendsto_eventually eventually_sequentiallyI [of m])
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3086
      qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3087
    qed auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3088
    then show fai: "f absolutely_integrable_on (\<Union>m. g ` F m)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3089
      using absolutely_integrable_restrict_UNIV by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3090
    show "integral ((\<Union>x. g ` F x)) f = integral (\<Union>n. F n) ?D"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3091
    proof (rule LIMSEQ_unique)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3092
      show "(\<lambda>n. integral (?U n) ?D) \<longlonglongrightarrow> integral (\<Union>x. g ` F x) f"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3093
        unfolding fg_int [symmetric]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3094
      proof (rule integral_countable_UN [OF fai])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3095
        show "g ` F m \<in> sets lebesgue" for m
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3096
        proof (rule differentiable_image_in_sets_lebesgue [OF F_leb])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3097
          show "g differentiable_on F m"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3098
            by (meson der_g differentiableI UnionI differentiable_on_def differentiable_on_subset rangeI subsetI)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3099
        qed auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3100
      qed
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  3101
    qed (use DU in metis)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3102
  next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3103
    assume fs: "f absolutely_integrable_on (\<Union>x. g ` F x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3104
      and b: "b = integral ((\<Union>x. g ` F x)) f"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3105
    have gF_leb: "g ` F m \<in> sets lebesgue" for m
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3106
    proof (rule differentiable_image_in_sets_lebesgue [OF F_leb])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3107
      show "g differentiable_on F m"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3108
        using der_g unfolding differentiable_def differentiable_on_def
72445
2c2de074832e tidying and removal of legacy name
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3109
        by (meson Sup_upper UNIV_I UnionI has_derivative_subset image_eqI)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3110
    qed auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3111
    have fgU: "\<And>n. f absolutely_integrable_on (\<Union>m\<le>n. g ` F m)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3112
      "(\<lambda>n. integral (\<Union>m\<le>n. g ` F m) f) \<longlonglongrightarrow> integral (\<Union>m. g ` F m) f"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3113
      using integral_countable_UN [OF fs gF_leb] by auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3114
    with iff have DUn: "?D absolutely_integrable_on ?U n"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3115
      and D_int: "integral (?U n) ?D = integral (\<Union>m\<le>n. g ` F m) f" for n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3116
      by (auto simp: image_UN)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3117
    let ?h = "\<lambda>x. if x \<in> (\<Union>n. F n) then norm(?D x) else 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3118
    have "(\<lambda>x. if x \<in> (\<Union>n. F n) then ?D x else 0) absolutely_integrable_on UNIV"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3119
    proof (rule dominated_convergence_absolutely_integrable)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3120
      show "(\<lambda>x. if x \<in> ?U k then ?D x else 0) absolutely_integrable_on UNIV" for k
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3121
        unfolding absolutely_integrable_restrict_UNIV using DUn by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3122
      let ?nD = "\<lambda>n x. if x \<in> ?U n then norm(?D x) else 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3123
      show "?h integrable_on UNIV"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3124
      proof (rule monotone_convergence_increasing [THEN conjunct1])
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3125
        show "?nD k integrable_on UNIV" for k
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3126
          using DUn
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3127
          unfolding integrable_restrict_UNIV absolutely_integrable_on_def by (simp add: image_UN)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3128
        { fix n::nat
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3129
          have "(norm \<circ> f) absolutely_integrable_on (\<Union>m\<le>n. g ` F m)"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  3130
            using absolutely_integrable_norm fgU by blast
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3131
          then have "integral (?U n) (norm \<circ> ?D) = integral (g ` ?U n) (norm \<circ> f)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3132
            using iff [of n "?lift \<circ> norm \<circ> f" "integral (g ` ?U n) (?lift \<circ> norm \<circ> f)"]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3133
            unfolding absolutely_integrable_on_1_iff integral_on_1_eq image_UN by (auto simp: o_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3134
        }
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3135
        moreover have "bounded (range (\<lambda>k. integral (g ` ?U k) (norm \<circ> f)))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3136
          unfolding bounded_iff
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3137
        proof (rule exI, clarify)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3138
          fix k
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3139
          show "norm (integral (g ` ?U k) (norm \<circ> f)) \<le> integral (g ` (\<Union>n. F n)) (norm \<circ> f)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3140
            unfolding integral_restrict_UNIV [of _ "norm \<circ> f", symmetric]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3141
          proof (rule integral_norm_bound_integral)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3142
            show "(\<lambda>x. if x \<in> g ` ?U k then (norm \<circ> f) x else 0) integrable_on UNIV"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3143
                 "(\<lambda>x. if x \<in> g ` (\<Union>n. F n) then (norm \<circ> f) x else 0) integrable_on UNIV"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3144
              using fgU fs
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3145
              unfolding absolutely_integrable_on_def o_def integrable_restrict_UNIV
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3146
              by (auto simp: image_UN)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3147
          qed auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3148
        qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3149
        ultimately show "bounded (range (\<lambda>k. integral UNIV (?nD k)))"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3150
          unfolding integral_restrict_UNIV image_UN [symmetric] o_def by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3151
      next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3152
        show "(\<lambda>k. if x \<in> ?U k then norm (?D x) else 0) \<longlonglongrightarrow> (if x \<in> (\<Union>n. F n) then norm (?D x) else 0)" for x
70365
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70139
diff changeset
  3153
          by (force intro: tendsto_eventually eventually_sequentiallyI)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3154
      qed auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3155
    next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3156
      show "(\<lambda>k. if x \<in> ?U k then ?D x else 0) \<longlonglongrightarrow> (if x \<in> (\<Union>n. F n) then ?D x else 0)" for x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3157
      proof clarsimp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3158
        fix n
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3159
        assume "x \<in> F n"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3160
        show "(\<lambda>m. if \<exists>j\<in>{..m}. x \<in> F j then ?D x else 0) \<longlonglongrightarrow> ?D x"
70365
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70139
diff changeset
  3161
          using \<open>x \<in> F n\<close> by (auto intro!: tendsto_eventually eventually_sequentiallyI [of n])
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3162
      qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3163
    qed auto
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3164
    then show Dai: "?D absolutely_integrable_on (\<Union>n. F n)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3165
      unfolding absolutely_integrable_restrict_UNIV by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3166
    show "integral (\<Union>n. F n) ?D = integral ((\<Union>x. g ` F x)) f"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3167
    proof (rule LIMSEQ_unique)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3168
      show "(\<lambda>n. integral (\<Union>m\<le>n. g ` F m) f) \<longlonglongrightarrow> integral (\<Union>n. F n) ?D"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3169
        unfolding D_int [symmetric] by (rule integral_countable_UN [OF Dai F_leb])
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 75462
diff changeset
  3170
    qed (use fgU in metis)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3171
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3172
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3173
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3174
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  3175
theorem has_absolute_integral_change_of_variables:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3176
  fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3177
  assumes S: "S \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3178
    and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3179
    and inj: "inj_on g S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3180
  shows "(\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3181
           integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) = b
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3182
     \<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b"
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  3183
proof -
70380
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  3184
  obtain C N where "fsigma C" and N: "N \<in> null_sets lebesgue" and CNS: "C \<union> N = S" and "disjnt C N"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3185
    using lebesgue_set_almost_fsigma [OF S] .
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3186
  then obtain F :: "nat \<Rightarrow> (real^'m::_) set"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3187
    where F: "range F \<subseteq> Collect compact" and Ceq: "C = Union(range F)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3188
    using fsigma_Union_compact by metis
70380
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  3189
  have "negligible N"
2b0dca68c3ee More analysis / measure theory material
paulson <lp15@cam.ac.uk>
parents: 70378
diff changeset
  3190
    using N by (simp add: negligible_iff_null_sets)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3191
  let ?D = "\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f (g x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3192
  have "?D absolutely_integrable_on C \<and> integral C ?D = b
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3193
    \<longleftrightarrow> f absolutely_integrable_on (g ` C) \<and> integral (g ` C) f = b"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3194
    unfolding Ceq
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3195
  proof (rule has_absolute_integral_change_of_variables_compact_family)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3196
    fix n x
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69272
diff changeset
  3197
    assume "x \<in> \<Union>(F ` UNIV)"
b021008c5397 removed legacy input syntax
haftmann
parents: 69272
diff changeset
  3198
    then show "(g has_derivative g' x) (at x within \<Union>(F ` UNIV))"
72445
2c2de074832e tidying and removal of legacy name
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3199
      using Ceq \<open>C \<union> N = S\<close> der_g has_derivative_subset by blast
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3200
  next
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69272
diff changeset
  3201
    have "\<Union>(F ` UNIV) \<subseteq> S"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3202
      using Ceq \<open>C \<union> N = S\<close> by blast
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69272
diff changeset
  3203
    then show "inj_on g (\<Union>(F ` UNIV))"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3204
      using inj by (meson inj_on_subset)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3205
  qed (use F in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3206
  moreover
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3207
  have "?D absolutely_integrable_on C \<and> integral C ?D = b
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3208
    \<longleftrightarrow> ?D absolutely_integrable_on S \<and> integral S ?D = b"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3209
  proof (rule conj_cong)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3210
    have neg: "negligible {x \<in> C - S. ?D x \<noteq> 0}" "negligible {x \<in> S - C. ?D x \<noteq> 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3211
      using CNS by (blast intro: negligible_subset [OF \<open>negligible N\<close>])+
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3212
    then show "(?D absolutely_integrable_on C) = (?D absolutely_integrable_on S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3213
      by (rule absolutely_integrable_spike_set_eq)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3214
    show "(integral C ?D = b) \<longleftrightarrow> (integral S ?D = b)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3215
      using integral_spike_set [OF neg] by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3216
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3217
  moreover
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3218
  have "f absolutely_integrable_on (g ` C) \<and> integral (g ` C) f = b
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3219
    \<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3220
  proof (rule conj_cong)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3221
    have "g differentiable_on N"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3222
      by (metis CNS der_g differentiable_def differentiable_on_def differentiable_on_subset sup.cobounded2)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3223
    with \<open>negligible N\<close>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3224
    have neg_gN: "negligible (g ` N)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3225
      by (blast intro: negligible_differentiable_image_negligible)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3226
    have neg: "negligible {x \<in> g ` C - g ` S. f x \<noteq> 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3227
              "negligible {x \<in> g ` S - g ` C. f x \<noteq> 0}"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3228
      using CNS by (blast intro: negligible_subset [OF neg_gN])+
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3229
    then show "(f absolutely_integrable_on g ` C) = (f absolutely_integrable_on g ` S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3230
      by (rule absolutely_integrable_spike_set_eq)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3231
    show "(integral (g ` C) f = b) \<longleftrightarrow> (integral (g ` S) f = b)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3232
      using integral_spike_set [OF neg] by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3233
  qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3234
  ultimately show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3235
    by simp
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3236
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3237
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3238
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  3239
corollary absolutely_integrable_change_of_variables:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3240
  fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3241
  assumes "S \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3242
    and "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3243
    and "inj_on g S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3244
  shows "f absolutely_integrable_on (g ` S)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3245
     \<longleftrightarrow> (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S"
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
  3246
  using assms has_absolute_integral_change_of_variables by blast
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3247
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  3248
corollary integral_change_of_variables:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3249
  fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3250
  assumes S: "S \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3251
    and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative g' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3252
    and inj: "inj_on g S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3253
    and disj: "(f absolutely_integrable_on (g ` S) \<or>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3254
        (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3255
  shows "integral (g ` S) f = integral S (\<lambda>x. \<bar>det (matrix (g' x))\<bar> *\<^sub>R f(g x))"
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
  3256
  using has_absolute_integral_change_of_variables [OF S der_g inj] disj
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
  3257
  by blast
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3258
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  3259
lemma has_absolute_integral_change_of_variables_1:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3260
  fixes f :: "real \<Rightarrow> real^'n::{finite,wellorder}" and g :: "real \<Rightarrow> real"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3261
  assumes S: "S \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3262
    and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_vector_derivative g' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3263
    and inj: "inj_on g S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3264
  shows "(\<lambda>x. \<bar>g' x\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3265
           integral S (\<lambda>x. \<bar>g' x\<bar> *\<^sub>R f(g x)) = b
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3266
     \<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3267
proof -
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3268
  let ?lift = "vec :: real \<Rightarrow> real^1"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3269
  let ?drop = "(\<lambda>x::real^1. x $ 1)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3270
  have S': "?lift ` S \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3271
    by (auto intro: differentiable_image_in_sets_lebesgue [OF S] differentiable_vec)
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
  3272
  have "((\<lambda>x. vec (g (x $ 1))) has_derivative (*\<^sub>R) (g' z)) (at (vec z) within ?lift ` S)"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3273
    if "z \<in> S" for z
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3274
    using der_g [OF that]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3275
    by (simp add: has_vector_derivative_def has_derivative_vector_1)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3276
  then have der': "\<And>x. x \<in> ?lift ` S \<Longrightarrow>
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
  3277
        (?lift \<circ> g \<circ> ?drop has_derivative (*\<^sub>R) (g' (?drop x))) (at x within ?lift ` S)"
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3278
    by (auto simp: o_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3279
  have inj': "inj_on (vec \<circ> g \<circ> ?drop) (vec ` S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3280
    using inj by (simp add: inj_on_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3281
  let ?fg = "\<lambda>x. \<bar>g' x\<bar> *\<^sub>R f(g x)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3282
  have "((\<lambda>x. ?fg x $ i) absolutely_integrable_on S \<and> ((\<lambda>x. ?fg x $ i) has_integral b $ i) S
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3283
    \<longleftrightarrow> (\<lambda>x. f x $ i) absolutely_integrable_on g ` S \<and> ((\<lambda>x. f x $ i) has_integral b $ i) (g ` S))" for i
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3284
    using has_absolute_integral_change_of_variables [OF S' der' inj', of "\<lambda>x. ?lift(f (?drop x) $ i)" "?lift (b$i)"]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3285
    unfolding integrable_on_1_iff integral_on_1_eq absolutely_integrable_on_1_iff absolutely_integrable_drop absolutely_integrable_on_def
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3286
    by (auto simp: image_comp o_def integral_vec1_eq has_integral_iff)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3287
  then have "?fg absolutely_integrable_on S \<and> (?fg has_integral b) S
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3288
         \<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> (f has_integral b) (g ` S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3289
    unfolding has_integral_componentwise_iff [where y=b]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3290
           absolutely_integrable_componentwise_iff [where f=f]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3291
           absolutely_integrable_componentwise_iff [where f = ?fg]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3292
    by (force simp: Basis_vec_def cart_eq_inner_axis)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3293
  then show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3294
    using absolutely_integrable_on_def by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3295
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3296
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3297
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  3298
corollary absolutely_integrable_change_of_variables_1:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3299
  fixes f :: "real \<Rightarrow> real^'n::{finite,wellorder}" and g :: "real \<Rightarrow> real"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3300
  assumes S: "S \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3301
    and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_vector_derivative g' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3302
    and inj: "inj_on g S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3303
  shows "(f absolutely_integrable_on g ` S \<longleftrightarrow>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3304
             (\<lambda>x. \<bar>g' x\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S)"
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
  3305
  using has_absolute_integral_change_of_variables_1 [OF assms] by auto
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3306
73928
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73648
diff changeset
  3307
text \<open>when @{term "n=1"}\<close>
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73648
diff changeset
  3308
lemma has_absolute_integral_change_of_variables_1':
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73648
diff changeset
  3309
  fixes f :: "real \<Rightarrow> real" and g :: "real \<Rightarrow> real"
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73648
diff changeset
  3310
  assumes S: "S \<in> sets lebesgue"
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73648
diff changeset
  3311
    and der_g: "\<And>x. x \<in> S \<Longrightarrow> (g has_field_derivative g' x) (at x within S)"
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73648
diff changeset
  3312
    and inj: "inj_on g S"
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73648
diff changeset
  3313
  shows "(\<lambda>x. \<bar>g' x\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and>
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73648
diff changeset
  3314
           integral S (\<lambda>x. \<bar>g' x\<bar> *\<^sub>R f(g x)) = b
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73648
diff changeset
  3315
     \<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b"
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73648
diff changeset
  3316
proof -
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73648
diff changeset
  3317
  have "(\<lambda>x. \<bar>g' x\<bar> *\<^sub>R vec (f(g x)) :: real ^ 1) absolutely_integrable_on S \<and>
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73648
diff changeset
  3318
           integral S (\<lambda>x. \<bar>g' x\<bar> *\<^sub>R vec (f(g x))) = (vec b :: real ^ 1)
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73648
diff changeset
  3319
         \<longleftrightarrow> (\<lambda>x. vec (f x) :: real ^ 1) absolutely_integrable_on (g ` S) \<and>
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73648
diff changeset
  3320
           integral (g ` S) (\<lambda>x. vec (f x)) = (vec b :: real ^ 1)"
75462
7448423e5dba Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
paulson <lp15@cam.ac.uk>
parents: 73933
diff changeset
  3321
    using assms unfolding has_real_derivative_iff_has_vector_derivative
73928
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73648
diff changeset
  3322
    by (intro has_absolute_integral_change_of_variables_1 assms) auto
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73648
diff changeset
  3323
  thus ?thesis
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73648
diff changeset
  3324
    by (simp add: absolutely_integrable_on_1_iff integral_on_1_eq)
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73648
diff changeset
  3325
qed
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73648
diff changeset
  3326
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3327
69683
8b3458ca0762 subsection is always %important
immler
parents: 69679
diff changeset
  3328
subsection\<open>Change of variables for integrals: special case of linear function\<close>
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3329
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  3330
lemma has_absolute_integral_change_of_variables_linear:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3331
  fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3332
  assumes "linear g"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3333
  shows "(\<lambda>x. \<bar>det (matrix g)\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S \<and>
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3334
           integral S (\<lambda>x. \<bar>det (matrix g)\<bar> *\<^sub>R f(g x)) = b
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3335
     \<longleftrightarrow> f absolutely_integrable_on (g ` S) \<and> integral (g ` S) f = b"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3336
proof (cases "det(matrix g) = 0")
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3337
  case True
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3338
  then have "negligible(g ` S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3339
    using assms det_nz_iff_inj negligible_linear_singular_image by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3340
  with True show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3341
    by (auto simp: absolutely_integrable_on_def integrable_negligible integral_negligible)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3342
next
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3343
  case False
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3344
  then obtain h where h: "\<And>x. x \<in> S \<Longrightarrow> h (g x) = x" "linear h"
71189
954ee5acaae0 Split off new HOL-Complex_Analysis session from HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71172
diff changeset
  3345
    using assms det_nz_iff_inj linear_injective_isomorphism by metis
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3346
  show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3347
  proof (rule has_absolute_integral_change_of_variables_invertible)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3348
    show "(g has_derivative g) (at x within S)" for x
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3349
      by (simp add: assms linear_imp_has_derivative)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3350
    show "continuous_on (g ` S) h"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3351
      using continuous_on_eq_continuous_within has_derivative_continuous linear_imp_has_derivative h by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3352
  qed (use h in auto)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3353
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3354
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  3355
lemma absolutely_integrable_change_of_variables_linear:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3356
  fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3357
  assumes "linear g"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3358
  shows "(\<lambda>x. \<bar>det (matrix g)\<bar> *\<^sub>R f(g x)) absolutely_integrable_on S
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3359
     \<longleftrightarrow> f absolutely_integrable_on (g ` S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3360
  using assms has_absolute_integral_change_of_variables_linear by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3361
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  3362
lemma absolutely_integrable_on_linear_image:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3363
  fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3364
  assumes "linear g"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3365
  shows "f absolutely_integrable_on (g ` S)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3366
     \<longleftrightarrow> (f \<circ> g) absolutely_integrable_on S \<or> det(matrix g) = 0"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3367
  unfolding assms absolutely_integrable_change_of_variables_linear [OF assms, symmetric] absolutely_integrable_on_scaleR_iff
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3368
  by (auto simp: set_integrable_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3369
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  3370
lemma integral_change_of_variables_linear:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3371
  fixes f :: "real^'m::{finite,wellorder} \<Rightarrow> real^'n" and g :: "real^'m::_ \<Rightarrow> real^'m::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3372
  assumes "linear g"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3373
      and "f absolutely_integrable_on (g ` S) \<or> (f \<circ> g) absolutely_integrable_on S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3374
    shows "integral (g ` S) f = \<bar>det (matrix g)\<bar> *\<^sub>R integral S (f \<circ> g)"
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  3375
proof -
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3376
  have "((\<lambda>x. \<bar>det (matrix g)\<bar> *\<^sub>R f (g x)) absolutely_integrable_on S) \<or> (f absolutely_integrable_on g ` S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3377
    using absolutely_integrable_on_linear_image assms by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3378
  moreover
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3379
  have ?thesis if "((\<lambda>x. \<bar>det (matrix g)\<bar> *\<^sub>R f (g x)) absolutely_integrable_on S)" "(f absolutely_integrable_on g ` S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3380
    using has_absolute_integral_change_of_variables_linear [OF \<open>linear g\<close>] that
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3381
    by (auto simp: o_def)
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3382
  ultimately show ?thesis
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3383
    using absolutely_integrable_change_of_variables_linear [OF \<open>linear g\<close>]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3384
    by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3385
qed
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3386
69683
8b3458ca0762 subsection is always %important
immler
parents: 69679
diff changeset
  3387
subsection\<open>Change of variable for measure\<close>
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3388
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  3389
lemma has_measure_differentiable_image:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3390
  fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3391
  assumes "S \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3392
      and "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3393
      and "inj_on f S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3394
  shows "f ` S \<in> lmeasurable \<and> measure lebesgue (f ` S) = m
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3395
     \<longleftrightarrow> ((\<lambda>x. \<bar>det (matrix (f' x))\<bar>) has_integral m) S"
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  3396
  using has_absolute_integral_change_of_variables [OF assms, of "\<lambda>x. (1::real^1)" "vec m"]
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  3397
  unfolding absolutely_integrable_on_1_iff integral_on_1_eq integrable_on_1_iff absolutely_integrable_on_def
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  3398
  by (auto simp: has_integral_iff lmeasurable_iff_integrable_on lmeasure_integral)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3399
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  3400
lemma measurable_differentiable_image_eq:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3401
  fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3402
  assumes "S \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3403
      and "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3404
      and "inj_on f S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3405
  shows "f ` S \<in> lmeasurable \<longleftrightarrow> (\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3406
  using has_measure_differentiable_image [OF assms]
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3407
  by blast
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3408
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  3409
lemma measurable_differentiable_image_alt:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3410
  fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3411
  assumes "S \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3412
    and "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3413
    and "inj_on f S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3414
  shows "f ` S \<in> lmeasurable \<longleftrightarrow> (\<lambda>x. \<bar>det (matrix (f' x))\<bar>) absolutely_integrable_on S"
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  3415
  using measurable_differentiable_image_eq [OF assms]
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  3416
  by (simp only: absolutely_integrable_on_iff_nonneg)
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3417
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  3418
lemma measure_differentiable_image_eq:
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3419
  fixes f :: "real^'n::{finite,wellorder} \<Rightarrow> real^'n::_"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3420
  assumes S: "S \<in> sets lebesgue"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3421
    and der_f: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3422
    and inj: "inj_on f S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3423
    and intS: "(\<lambda>x. \<bar>det (matrix (f' x))\<bar>) integrable_on S"
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3424
  shows "measure lebesgue (f ` S) = integral S (\<lambda>x. \<bar>det (matrix (f' x))\<bar>)"
69720
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  3425
  using measurable_differentiable_image_eq [OF S der_f inj]
be6634e99e09 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  3426
        assms has_measure_differentiable_image by blast
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3427
82518
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79945
diff changeset
  3428
lemma has_absolute_integral_reflect_real:
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79945
diff changeset
  3429
  fixes f :: "real \<Rightarrow> real"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79945
diff changeset
  3430
  assumes "uminus ` A \<subseteq> B" "uminus ` B \<subseteq> A" "A \<in> sets lebesgue"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79945
diff changeset
  3431
  shows "(\<lambda>x. f (-x)) absolutely_integrable_on A \<and> integral A (\<lambda>x. f (-x)) = b \<longleftrightarrow>
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79945
diff changeset
  3432
         f absolutely_integrable_on B \<and> integral B f = b"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79945
diff changeset
  3433
proof -
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79945
diff changeset
  3434
  have bij: "bij_betw uminus A B"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79945
diff changeset
  3435
    by (rule bij_betwI[of _ _ _ uminus]) (use assms(1,2) in auto)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79945
diff changeset
  3436
  have "((\<lambda>x. \<bar>-1\<bar> *\<^sub>R f (-x)) absolutely_integrable_on A \<and>
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79945
diff changeset
  3437
          integral A (\<lambda>x. \<bar>-1\<bar> *\<^sub>R f (-x)) = b) \<longleftrightarrow>
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79945
diff changeset
  3438
        (f absolutely_integrable_on uminus ` A \<and>
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79945
diff changeset
  3439
          integral (uminus ` A) f = b)" using assms
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79945
diff changeset
  3440
    by (intro has_absolute_integral_change_of_variables_1') (auto intro!: derivative_eq_intros)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79945
diff changeset
  3441
  also have "uminus ` A = B"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79945
diff changeset
  3442
    using bij by (simp add: bij_betw_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79945
diff changeset
  3443
  finally show ?thesis
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79945
diff changeset
  3444
    by simp
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79945
diff changeset
  3445
qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79945
diff changeset
  3446
67998
73a5a33486ee Change of variables proof
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3447
end