src/HOL/Analysis/Tagged_Division.thy
author paulson
Wed, 02 Aug 2017 23:15:15 +0200
changeset 66314 52914a618299
parent 66306 13b051ebc6c5
child 66317 a9bb833ee971
permissions -rw-r--r--
removed all "guess"
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
     1
(*  Title:      HOL/Analysis/Tagged_Division.thy
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
     2
    Author:     John Harrison
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
     3
    Author:     Robert Himmelmann, TU Muenchen (Translation from HOL light); proofs reworked by LCP
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
     4
*)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
     5
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
     6
section \<open>Tagged divisions used for the Henstock-Kurzweil gauge integration\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
     7
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
     8
theory Tagged_Division
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
     9
imports
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    10
  Topology_Euclidean_Space
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    11
begin
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    12
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
    13
lemma finite_product_dependent: (*FIXME DELETE*)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    14
  assumes "finite s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    15
    and "\<And>x. x \<in> s \<Longrightarrow> finite (t x)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    16
  shows "finite {(i, j) |i j. i \<in> s \<and> j \<in> t i}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    17
  using assms
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    18
proof induct
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    19
  case (insert x s)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    20
  have *: "{(i, j) |i j. i \<in> insert x s \<and> j \<in> t i} =
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    21
    (\<lambda>y. (x,y)) ` (t x) \<union> {(i, j) |i j. i \<in> s \<and> j \<in> t i}" by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    22
  show ?case
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    23
    unfolding *
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    24
    apply (rule finite_UnI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    25
    using insert
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    26
    apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    27
    done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    28
qed auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    29
66294
0442b3f45556 refactored some HORRIBLE integration proofs
paulson <lp15@cam.ac.uk>
parents: 66193
diff changeset
    30
lemma sum_Sigma_product:
0442b3f45556 refactored some HORRIBLE integration proofs
paulson <lp15@cam.ac.uk>
parents: 66193
diff changeset
    31
  assumes "finite S"
0442b3f45556 refactored some HORRIBLE integration proofs
paulson <lp15@cam.ac.uk>
parents: 66193
diff changeset
    32
    and "\<And>i. i \<in> S \<Longrightarrow> finite (T i)"
0442b3f45556 refactored some HORRIBLE integration proofs
paulson <lp15@cam.ac.uk>
parents: 66193
diff changeset
    33
  shows "(\<Sum>i\<in>S. sum (x i) (T i)) = (\<Sum>(i, j)\<in>Sigma S T. x i j)"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    34
  using assms
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    35
proof induct
66294
0442b3f45556 refactored some HORRIBLE integration proofs
paulson <lp15@cam.ac.uk>
parents: 66193
diff changeset
    36
  case empty
0442b3f45556 refactored some HORRIBLE integration proofs
paulson <lp15@cam.ac.uk>
parents: 66193
diff changeset
    37
  then show ?case
0442b3f45556 refactored some HORRIBLE integration proofs
paulson <lp15@cam.ac.uk>
parents: 66193
diff changeset
    38
    by simp
0442b3f45556 refactored some HORRIBLE integration proofs
paulson <lp15@cam.ac.uk>
parents: 66193
diff changeset
    39
next
0442b3f45556 refactored some HORRIBLE integration proofs
paulson <lp15@cam.ac.uk>
parents: 66193
diff changeset
    40
  case (insert a S)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    41
  show ?case
66294
0442b3f45556 refactored some HORRIBLE integration proofs
paulson <lp15@cam.ac.uk>
parents: 66193
diff changeset
    42
    by (simp add: insert.hyps(1) insert.prems sum.Sigma)
0442b3f45556 refactored some HORRIBLE integration proofs
paulson <lp15@cam.ac.uk>
parents: 66193
diff changeset
    43
qed
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    44
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    45
lemmas scaleR_simps = scaleR_zero_left scaleR_minus_left scaleR_left_diff_distrib
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    46
  scaleR_zero_right scaleR_minus_right scaleR_right_diff_distrib scaleR_eq_0_iff
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    47
  scaleR_cancel_left scaleR_cancel_right scaleR_add_right scaleR_add_left real_vector_class.scaleR_one
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    48
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    49
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    50
subsection \<open>Sundries\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    51
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    52
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    53
text\<open>A transitive relation is well-founded if all initial segments are finite.\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    54
lemma wf_finite_segments:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    55
  assumes "irrefl r" and "trans r" and "\<And>x. finite {y. (y, x) \<in> r}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    56
  shows "wf (r)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    57
  apply (simp add: trans_wf_iff wf_iff_acyclic_if_finite converse_def assms)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    58
  using acyclic_def assms irrefl_def trans_Restr by fastforce
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    59
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    60
text\<open>For creating values between @{term u} and @{term v}.\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    61
lemma scaling_mono:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    62
  fixes u::"'a::linordered_field"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    63
  assumes "u \<le> v" "0 \<le> r" "r \<le> s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    64
    shows "u + r * (v - u) / s \<le> v"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    65
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    66
  have "r/s \<le> 1" using assms
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    67
    using divide_le_eq_1 by fastforce
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    68
  then have "(r/s) * (v - u) \<le> 1 * (v - u)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    69
    by (meson diff_ge_0_iff_ge mult_right_mono \<open>u \<le> v\<close>)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    70
  then show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    71
    by (simp add: field_simps)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    72
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    73
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    74
subsection \<open>Some useful lemmas about intervals.\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    75
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    76
lemma interior_subset_union_intervals:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    77
  assumes "i = cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    78
    and "j = cbox c d"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    79
    and "interior j \<noteq> {}"
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
    80
    and "i \<subseteq> j \<union> S"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    81
    and "interior i \<inter> interior j = {}"
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
    82
  shows "interior i \<subseteq> interior S"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    83
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    84
  have "box a b \<inter> cbox c d = {}"
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
    85
     using inter_interval_mixed_eq_empty[of c d a b] assms
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
    86
     unfolding interior_cbox by auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    87
  moreover
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
    88
  have "box a b \<subseteq> cbox c d \<union> S"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    89
    apply (rule order_trans,rule box_subset_cbox)
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
    90
    using assms by auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    91
  ultimately
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    92
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    93
    unfolding assms interior_cbox
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    94
      by auto (metis IntI UnE empty_iff interior_maximal open_box subsetCE subsetI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    95
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    96
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    97
lemma interior_Union_subset_cbox:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    98
  assumes "finite f"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    99
  assumes f: "\<And>s. s \<in> f \<Longrightarrow> \<exists>a b. s = cbox a b" "\<And>s. s \<in> f \<Longrightarrow> interior s \<subseteq> t"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   100
    and t: "closed t"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   101
  shows "interior (\<Union>f) \<subseteq> t"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   102
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   103
  have [simp]: "s \<in> f \<Longrightarrow> closed s" for s
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   104
    using f by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   105
  define E where "E = {s\<in>f. interior s = {}}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   106
  then have "finite E" "E \<subseteq> {s\<in>f. interior s = {}}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   107
    using \<open>finite f\<close> by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   108
  then have "interior (\<Union>f) = interior (\<Union>(f - E))"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   109
  proof (induction E rule: finite_subset_induct')
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   110
    case (insert s f')
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   111
    have "interior (\<Union>(f - insert s f') \<union> s) = interior (\<Union>(f - insert s f'))"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   112
      using insert.hyps \<open>finite f\<close> by (intro interior_closed_Un_empty_interior) auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   113
    also have "\<Union>(f - insert s f') \<union> s = \<Union>(f - f')"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   114
      using insert.hyps by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   115
    finally show ?case
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   116
      by (simp add: insert.IH)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   117
  qed simp
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   118
  also have "\<dots> \<subseteq> \<Union>(f - E)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   119
    by (rule interior_subset)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   120
  also have "\<dots> \<subseteq> t"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   121
  proof (rule Union_least)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   122
    fix s assume "s \<in> f - E"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   123
    with f[of s] obtain a b where s: "s \<in> f" "s = cbox a b" "box a b \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   124
      by (fastforce simp: E_def)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   125
    have "closure (interior s) \<subseteq> closure t"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   126
      by (intro closure_mono f \<open>s \<in> f\<close>)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   127
    with s \<open>closed t\<close> show "s \<subseteq> t"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   128
      by simp
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   129
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   130
  finally show ?thesis .
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   131
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   132
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   133
lemma Int_interior_Union_intervals:
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   134
    "\<lbrakk>finite \<F>; open S; \<And>T. T\<in>\<F> \<Longrightarrow> \<exists>a b. T = cbox a b; \<And>T. T\<in>\<F> \<Longrightarrow> S \<inter> (interior T) = {}\<rbrakk> 
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   135
    \<Longrightarrow> S \<inter> interior (\<Union>\<F>) = {}"
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   136
  using interior_Union_subset_cbox[of \<F> "UNIV - S"] by auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   137
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   138
lemma interval_split:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   139
  fixes a :: "'a::euclidean_space"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   140
  assumes "k \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   141
  shows
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   142
    "cbox a b \<inter> {x. x\<bullet>k \<le> c} = cbox a (\<Sum>i\<in>Basis. (if i = k then min (b\<bullet>k) c else b\<bullet>i) *\<^sub>R i)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   143
    "cbox a b \<inter> {x. x\<bullet>k \<ge> c} = cbox (\<Sum>i\<in>Basis. (if i = k then max (a\<bullet>k) c else a\<bullet>i) *\<^sub>R i) b"
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   144
  using assms by (rule_tac set_eqI; auto simp: mem_box)+
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   145
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   146
lemma interval_not_empty: "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow> cbox a b \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   147
  by (simp add: box_ne_empty)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   148
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   149
subsection \<open>Bounds on intervals where they exist.\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   150
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   151
definition interval_upperbound :: "('a::euclidean_space) set \<Rightarrow> 'a"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   152
  where "interval_upperbound s = (\<Sum>i\<in>Basis. (SUP x:s. x\<bullet>i) *\<^sub>R i)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   153
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   154
definition interval_lowerbound :: "('a::euclidean_space) set \<Rightarrow> 'a"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   155
  where "interval_lowerbound s = (\<Sum>i\<in>Basis. (INF x:s. x\<bullet>i) *\<^sub>R i)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   156
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   157
lemma interval_upperbound[simp]:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   158
  "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   159
    interval_upperbound (cbox a b) = (b::'a::euclidean_space)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
   160
  unfolding interval_upperbound_def euclidean_representation_sum cbox_def
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   161
  by (safe intro!: cSup_eq) auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   162
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   163
lemma interval_lowerbound[simp]:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   164
  "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   165
    interval_lowerbound (cbox a b) = (a::'a::euclidean_space)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
   166
  unfolding interval_lowerbound_def euclidean_representation_sum cbox_def
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   167
  by (safe intro!: cInf_eq) auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   168
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   169
lemmas interval_bounds = interval_upperbound interval_lowerbound
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   170
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   171
lemma
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   172
  fixes X::"real set"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   173
  shows interval_upperbound_real[simp]: "interval_upperbound X = Sup X"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   174
    and interval_lowerbound_real[simp]: "interval_lowerbound X = Inf X"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   175
  by (auto simp: interval_upperbound_def interval_lowerbound_def)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   176
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   177
lemma interval_bounds'[simp]:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   178
  assumes "cbox a b \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   179
  shows "interval_upperbound (cbox a b) = b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   180
    and "interval_lowerbound (cbox a b) = a"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   181
  using assms unfolding box_ne_empty by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   182
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   183
lemma interval_upperbound_Times:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   184
  assumes "A \<noteq> {}" and "B \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   185
  shows "interval_upperbound (A \<times> B) = (interval_upperbound A, interval_upperbound B)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   186
proof-
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   187
  from assms have fst_image_times': "A = fst ` (A \<times> B)" by simp
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   188
  have "(\<Sum>i\<in>Basis. (SUP x:A \<times> B. x \<bullet> (i, 0)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (SUP x:A. x \<bullet> i) *\<^sub>R i)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   189
      by (subst (2) fst_image_times') (simp del: fst_image_times add: o_def inner_Pair_0)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   190
  moreover from assms have snd_image_times': "B = snd ` (A \<times> B)" by simp
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   191
  have "(\<Sum>i\<in>Basis. (SUP x:A \<times> B. x \<bullet> (0, i)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (SUP x:B. x \<bullet> i) *\<^sub>R i)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   192
      by (subst (2) snd_image_times') (simp del: snd_image_times add: o_def inner_Pair_0)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   193
  ultimately show ?thesis unfolding interval_upperbound_def
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
   194
      by (subst sum_Basis_prod_eq) (auto simp add: sum_prod)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   195
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   196
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   197
lemma interval_lowerbound_Times:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   198
  assumes "A \<noteq> {}" and "B \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   199
  shows "interval_lowerbound (A \<times> B) = (interval_lowerbound A, interval_lowerbound B)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   200
proof-
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   201
  from assms have fst_image_times': "A = fst ` (A \<times> B)" by simp
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   202
  have "(\<Sum>i\<in>Basis. (INF x:A \<times> B. x \<bullet> (i, 0)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (INF x:A. x \<bullet> i) *\<^sub>R i)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   203
      by (subst (2) fst_image_times') (simp del: fst_image_times add: o_def inner_Pair_0)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   204
  moreover from assms have snd_image_times': "B = snd ` (A \<times> B)" by simp
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   205
  have "(\<Sum>i\<in>Basis. (INF x:A \<times> B. x \<bullet> (0, i)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (INF x:B. x \<bullet> i) *\<^sub>R i)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   206
      by (subst (2) snd_image_times') (simp del: snd_image_times add: o_def inner_Pair_0)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   207
  ultimately show ?thesis unfolding interval_lowerbound_def
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
   208
      by (subst sum_Basis_prod_eq) (auto simp add: sum_prod)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   209
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   210
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   211
subsection \<open>The notion of a gauge --- simply an open set containing the point.\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   212
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   213
definition "gauge d \<longleftrightarrow> (\<forall>x. x \<in> d x \<and> open (d x))"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   214
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   215
lemma gaugeI:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   216
  assumes "\<And>x. x \<in> g x"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   217
    and "\<And>x. open (g x)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   218
  shows "gauge g"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   219
  using assms unfolding gauge_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   220
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   221
lemma gaugeD[dest]:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   222
  assumes "gauge d"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   223
  shows "x \<in> d x"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   224
    and "open (d x)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   225
  using assms unfolding gauge_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   226
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   227
lemma gauge_ball_dependent: "\<forall>x. 0 < e x \<Longrightarrow> gauge (\<lambda>x. ball x (e x))"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   228
  unfolding gauge_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   229
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   230
lemma gauge_ball[intro]: "0 < e \<Longrightarrow> gauge (\<lambda>x. ball x e)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   231
  unfolding gauge_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   232
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   233
lemma gauge_trivial[intro!]: "gauge (\<lambda>x. ball x 1)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   234
  by (rule gauge_ball) auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   235
66154
bc5e6461f759 Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents: 66113
diff changeset
   236
lemma gauge_Int[intro]: "gauge d1 \<Longrightarrow> gauge d2 \<Longrightarrow> gauge (\<lambda>x. d1 x \<inter> d2 x)"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   237
  unfolding gauge_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   238
66164
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 66154
diff changeset
   239
lemma gauge_reflect:
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   240
  fixes \<gamma> :: "'a::euclidean_space \<Rightarrow> 'a set"
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   241
  shows "gauge \<gamma> \<Longrightarrow> gauge (\<lambda>x. uminus ` \<gamma> (- x))"
66164
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 66154
diff changeset
   242
  using equation_minus_iff
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 66154
diff changeset
   243
  by (auto simp: gauge_def surj_def intro!: open_surjective_linear_image linear_uminus)
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 66154
diff changeset
   244
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 66154
diff changeset
   245
lemma gauge_Inter:
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   246
  assumes "finite S"
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   247
    and "\<And>d. d\<in>S \<Longrightarrow> gauge (f d)"
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   248
  shows "gauge (\<lambda>x. \<Inter>{f d x | d. d \<in> S})"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   249
proof -
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   250
  have *: "\<And>x. {f d x |d. d \<in> S} = (\<lambda>d. f d x) ` S"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   251
    by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   252
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   253
    unfolding gauge_def unfolding *
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   254
    using assms unfolding Ball_def Inter_iff mem_Collect_eq gauge_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   255
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   256
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   257
lemma gauge_existence_lemma:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   258
  "(\<forall>x. \<exists>d :: real. p x \<longrightarrow> 0 < d \<and> q d x) \<longleftrightarrow> (\<forall>x. \<exists>d>0. p x \<longrightarrow> q d x)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   259
  by (metis zero_less_one)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   260
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   261
subsection \<open>Attempt a systematic general set of "offset" results for components.\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   262
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   263
lemma gauge_modify:
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   264
  assumes "(\<forall>S. open S \<longrightarrow> open {x. f(x) \<in> S})" "gauge d"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   265
  shows "gauge (\<lambda>x. {y. f y \<in> d (f x)})"
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   266
  using assms unfolding gauge_def by force
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   267
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   268
subsection \<open>Divisions.\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   269
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   270
definition division_of (infixl "division'_of" 40)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   271
where
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   272
  "s division_of i \<longleftrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   273
    finite s \<and>
66113
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
   274
    (\<forall>K\<in>s. K \<subseteq> i \<and> K \<noteq> {} \<and> (\<exists>a b. K = cbox a b)) \<and>
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
   275
    (\<forall>K1\<in>s. \<forall>K2\<in>s. K1 \<noteq> K2 \<longrightarrow> interior(K1) \<inter> interior(K2) = {}) \<and>
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   276
    (\<Union>s = i)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   277
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   278
lemma division_ofD[dest]:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   279
  assumes "s division_of i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   280
  shows "finite s"
66113
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
   281
    and "\<And>K. K \<in> s \<Longrightarrow> K \<subseteq> i"
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
   282
    and "\<And>K. K \<in> s \<Longrightarrow> K \<noteq> {}"
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
   283
    and "\<And>K. K \<in> s \<Longrightarrow> \<exists>a b. K = cbox a b"
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
   284
    and "\<And>K1 K2. K1 \<in> s \<Longrightarrow> K2 \<in> s \<Longrightarrow> K1 \<noteq> K2 \<Longrightarrow> interior(K1) \<inter> interior(K2) = {}"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   285
    and "\<Union>s = i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   286
  using assms unfolding division_of_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   287
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   288
lemma division_ofI:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   289
  assumes "finite s"
66113
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
   290
    and "\<And>K. K \<in> s \<Longrightarrow> K \<subseteq> i"
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
   291
    and "\<And>K. K \<in> s \<Longrightarrow> K \<noteq> {}"
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
   292
    and "\<And>K. K \<in> s \<Longrightarrow> \<exists>a b. K = cbox a b"
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
   293
    and "\<And>K1 K2. K1 \<in> s \<Longrightarrow> K2 \<in> s \<Longrightarrow> K1 \<noteq> K2 \<Longrightarrow> interior K1 \<inter> interior K2 = {}"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   294
    and "\<Union>s = i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   295
  shows "s division_of i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   296
  using assms unfolding division_of_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   297
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   298
lemma division_of_finite: "s division_of i \<Longrightarrow> finite s"
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents: 66295
diff changeset
   299
  by auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   300
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   301
lemma division_of_self[intro]: "cbox a b \<noteq> {} \<Longrightarrow> {cbox a b} division_of (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   302
  unfolding division_of_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   303
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   304
lemma division_of_trivial[simp]: "s division_of {} \<longleftrightarrow> s = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   305
  unfolding division_of_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   306
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   307
lemma division_of_sing[simp]:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   308
  "s division_of cbox a (a::'a::euclidean_space) \<longleftrightarrow> s = {cbox a a}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   309
  (is "?l = ?r")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   310
proof
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   311
  assume ?r
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   312
  moreover
66113
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
   313
  { fix K
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
   314
    assume "s = {{a}}" "K\<in>s"
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
   315
    then have "\<exists>x y. K = cbox x y"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   316
      apply (rule_tac x=a in exI)+
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   317
      apply (force simp: cbox_sing)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   318
      done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   319
  }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   320
  ultimately show ?l
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   321
    unfolding division_of_def cbox_sing by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   322
next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   323
  assume ?l
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   324
  {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   325
    fix x
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   326
    assume x: "x \<in> s" have "x = {a}"
66193
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
   327
      by (metis \<open>s division_of cbox a a\<close> cbox_sing division_ofD(2) division_ofD(3) subset_singletonD x)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   328
  }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   329
  moreover have "s \<noteq> {}"
66193
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
   330
    using \<open>s division_of cbox a a\<close> by auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   331
  ultimately show ?r
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   332
    unfolding cbox_sing by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   333
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   334
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   335
lemma elementary_empty: obtains p where "p division_of {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   336
  unfolding division_of_trivial by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   337
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   338
lemma elementary_interval: obtains p where "p division_of (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   339
  by (metis division_of_trivial division_of_self)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   340
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   341
lemma division_contains: "s division_of i \<Longrightarrow> \<forall>x\<in>i. \<exists>k\<in>s. x \<in> k"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   342
  unfolding division_of_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   343
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   344
lemma forall_in_division:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   345
  "d division_of i \<Longrightarrow> (\<forall>x\<in>d. P x) \<longleftrightarrow> (\<forall>a b. cbox a b \<in> d \<longrightarrow> P (cbox a b))"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   346
  unfolding division_of_def by fastforce
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   347
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents: 66295
diff changeset
   348
lemma cbox_division_memE:
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents: 66295
diff changeset
   349
  assumes \<D>: "K \<in> \<D>" "\<D> division_of S"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents: 66295
diff changeset
   350
  obtains a b where "K = cbox a b" "K \<noteq> {}" "K \<subseteq> S"
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents: 66295
diff changeset
   351
  using assms unfolding division_of_def by metis
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents: 66295
diff changeset
   352
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   353
lemma division_of_subset:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   354
  assumes "p division_of (\<Union>p)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   355
    and "q \<subseteq> p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   356
  shows "q division_of (\<Union>q)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   357
proof (rule division_ofI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   358
  note * = division_ofD[OF assms(1)]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   359
  show "finite q"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   360
    using "*"(1) assms(2) infinite_super by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   361
  {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   362
    fix k
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   363
    assume "k \<in> q"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   364
    then have kp: "k \<in> p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   365
      using assms(2) by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   366
    show "k \<subseteq> \<Union>q"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   367
      using \<open>k \<in> q\<close> by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   368
    show "\<exists>a b. k = cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   369
      using *(4)[OF kp] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   370
    show "k \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   371
      using *(3)[OF kp] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   372
  }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   373
  fix k1 k2
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   374
  assume "k1 \<in> q" "k2 \<in> q" "k1 \<noteq> k2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   375
  then have **: "k1 \<in> p" "k2 \<in> p" "k1 \<noteq> k2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   376
    using assms(2) by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   377
  show "interior k1 \<inter> interior k2 = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   378
    using *(5)[OF **] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   379
qed auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   380
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   381
lemma division_of_union_self[intro]: "p division_of s \<Longrightarrow> p division_of (\<Union>p)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   382
  unfolding division_of_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   383
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   384
lemma division_inter:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   385
  fixes s1 s2 :: "'a::euclidean_space set"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   386
  assumes "p1 division_of s1"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   387
    and "p2 division_of s2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   388
  shows "{k1 \<inter> k2 | k1 k2. k1 \<in> p1 \<and> k2 \<in> p2 \<and> k1 \<inter> k2 \<noteq> {}} division_of (s1 \<inter> s2)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   389
  (is "?A' division_of _")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   390
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   391
  let ?A = "{s. s \<in>  (\<lambda>(k1,k2). k1 \<inter> k2) ` (p1 \<times> p2) \<and> s \<noteq> {}}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   392
  have *: "?A' = ?A" by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   393
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   394
    unfolding *
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   395
  proof (rule division_ofI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   396
    have "?A \<subseteq> (\<lambda>(x, y). x \<inter> y) ` (p1 \<times> p2)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   397
      by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   398
    moreover have "finite (p1 \<times> p2)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   399
      using assms unfolding division_of_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   400
    ultimately show "finite ?A" by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   401
    have *: "\<And>s. \<Union>{x\<in>s. x \<noteq> {}} = \<Union>s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   402
      by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   403
    show "\<Union>?A = s1 \<inter> s2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   404
      apply (rule set_eqI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   405
      unfolding * and UN_iff
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   406
      using division_ofD(6)[OF assms(1)] and division_ofD(6)[OF assms(2)]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   407
      apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   408
      done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   409
    {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   410
      fix k
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   411
      assume "k \<in> ?A"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   412
      then obtain k1 k2 where k: "k = k1 \<inter> k2" "k1 \<in> p1" "k2 \<in> p2" "k \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   413
        by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   414
      then show "k \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   415
        by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   416
      show "k \<subseteq> s1 \<inter> s2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   417
        using division_ofD(2)[OF assms(1) k(2)] and division_ofD(2)[OF assms(2) k(3)]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   418
        unfolding k by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   419
      obtain a1 b1 where k1: "k1 = cbox a1 b1"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   420
        using division_ofD(4)[OF assms(1) k(2)] by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   421
      obtain a2 b2 where k2: "k2 = cbox a2 b2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   422
        using division_ofD(4)[OF assms(2) k(3)] by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   423
      show "\<exists>a b. k = cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   424
        unfolding k k1 k2 unfolding Int_interval by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   425
    }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   426
    fix k1 k2
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   427
    assume "k1 \<in> ?A"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   428
    then obtain x1 y1 where k1: "k1 = x1 \<inter> y1" "x1 \<in> p1" "y1 \<in> p2" "k1 \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   429
      by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   430
    assume "k2 \<in> ?A"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   431
    then obtain x2 y2 where k2: "k2 = x2 \<inter> y2" "x2 \<in> p1" "y2 \<in> p2" "k2 \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   432
      by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   433
    assume "k1 \<noteq> k2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   434
    then have th: "x1 \<noteq> x2 \<or> y1 \<noteq> y2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   435
      unfolding k1 k2 by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   436
    have *: "interior x1 \<inter> interior x2 = {} \<or> interior y1 \<inter> interior y2 = {} \<Longrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   437
      interior (x1 \<inter> y1) \<subseteq> interior x1 \<Longrightarrow> interior (x1 \<inter> y1) \<subseteq> interior y1 \<Longrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   438
      interior (x2 \<inter> y2) \<subseteq> interior x2 \<Longrightarrow> interior (x2 \<inter> y2) \<subseteq> interior y2 \<Longrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   439
      interior (x1 \<inter> y1) \<inter> interior (x2 \<inter> y2) = {}" by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   440
    show "interior k1 \<inter> interior k2 = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   441
      unfolding k1 k2
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   442
      apply (rule *)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   443
      using assms division_ofD(5) k1 k2(2) k2(3) th apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   444
      done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   445
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   446
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   447
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   448
lemma division_inter_1:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   449
  assumes "d division_of i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   450
    and "cbox a (b::'a::euclidean_space) \<subseteq> i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   451
  shows "{cbox a b \<inter> k | k. k \<in> d \<and> cbox a b \<inter> k \<noteq> {}} division_of (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   452
proof (cases "cbox a b = {}")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   453
  case True
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   454
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   455
    unfolding True and division_of_trivial by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   456
next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   457
  case False
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   458
  have *: "cbox a b \<inter> i = cbox a b" using assms(2) by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   459
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   460
    using division_inter[OF division_of_self[OF False] assms(1)]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   461
    unfolding * by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   462
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   463
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   464
lemma elementary_inter:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   465
  fixes s t :: "'a::euclidean_space set"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   466
  assumes "p1 division_of s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   467
    and "p2 division_of t"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   468
  shows "\<exists>p. p division_of (s \<inter> t)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   469
using assms division_inter by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   470
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   471
lemma elementary_inters:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   472
  assumes "finite f"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   473
    and "f \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   474
    and "\<forall>s\<in>f. \<exists>p. p division_of (s::('a::euclidean_space) set)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   475
  shows "\<exists>p. p division_of (\<Inter>f)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   476
  using assms
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   477
proof (induct f rule: finite_induct)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   478
  case (insert x f)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   479
  show ?case
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   480
  proof (cases "f = {}")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   481
    case True
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   482
    then show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   483
      unfolding True using insert by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   484
  next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   485
    case False
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   486
    obtain p where "p division_of \<Inter>f"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   487
      using insert(3)[OF False insert(5)[unfolded ball_simps,THEN conjunct2]] ..
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   488
    moreover obtain px where "px division_of x"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   489
      using insert(5)[rule_format,OF insertI1] ..
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   490
    ultimately show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   491
      by (simp add: elementary_inter Inter_insert)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   492
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   493
qed auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   494
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   495
lemma division_disjoint_union:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   496
  assumes "p1 division_of s1"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   497
    and "p2 division_of s2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   498
    and "interior s1 \<inter> interior s2 = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   499
  shows "(p1 \<union> p2) division_of (s1 \<union> s2)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   500
proof (rule division_ofI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   501
  note d1 = division_ofD[OF assms(1)]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   502
  note d2 = division_ofD[OF assms(2)]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   503
  show "finite (p1 \<union> p2)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   504
    using d1(1) d2(1) by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   505
  show "\<Union>(p1 \<union> p2) = s1 \<union> s2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   506
    using d1(6) d2(6) by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   507
  {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   508
    fix k1 k2
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   509
    assume as: "k1 \<in> p1 \<union> p2" "k2 \<in> p1 \<union> p2" "k1 \<noteq> k2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   510
    moreover
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   511
    let ?g="interior k1 \<inter> interior k2 = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   512
    {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   513
      assume as: "k1\<in>p1" "k2\<in>p2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   514
      have ?g
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   515
        using interior_mono[OF d1(2)[OF as(1)]] interior_mono[OF d2(2)[OF as(2)]]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   516
        using assms(3) by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   517
    }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   518
    moreover
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   519
    {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   520
      assume as: "k1\<in>p2" "k2\<in>p1"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   521
      have ?g
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   522
        using interior_mono[OF d1(2)[OF as(2)]] interior_mono[OF d2(2)[OF as(1)]]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   523
        using assms(3) by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   524
    }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   525
    ultimately show ?g
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   526
      using d1(5)[OF _ _ as(3)] and d2(5)[OF _ _ as(3)] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   527
  }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   528
  fix k
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   529
  assume k: "k \<in> p1 \<union> p2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   530
  show "k \<subseteq> s1 \<union> s2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   531
    using k d1(2) d2(2) by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   532
  show "k \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   533
    using k d1(3) d2(3) by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   534
  show "\<exists>a b. k = cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   535
    using k d1(4) d2(4) by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   536
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   537
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   538
lemma partial_division_extend_1:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   539
  fixes a b c d :: "'a::euclidean_space"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   540
  assumes incl: "cbox c d \<subseteq> cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   541
    and nonempty: "cbox c d \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   542
  obtains p where "p division_of (cbox a b)" "cbox c d \<in> p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   543
proof
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   544
  let ?B = "\<lambda>f::'a\<Rightarrow>'a \<times> 'a.
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   545
    cbox (\<Sum>i\<in>Basis. (fst (f i) \<bullet> i) *\<^sub>R i) (\<Sum>i\<in>Basis. (snd (f i) \<bullet> i) *\<^sub>R i)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   546
  define p where "p = ?B ` (Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)})"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   547
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   548
  show "cbox c d \<in> p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   549
    unfolding p_def
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   550
    by (auto simp add: box_eq_empty cbox_def intro!: image_eqI[where x="\<lambda>(i::'a)\<in>Basis. (c, d)"])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   551
  {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   552
    fix i :: 'a
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   553
    assume "i \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   554
    with incl nonempty have "a \<bullet> i \<le> c \<bullet> i" "c \<bullet> i \<le> d \<bullet> i" "d \<bullet> i \<le> b \<bullet> i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   555
      unfolding box_eq_empty subset_box by (auto simp: not_le)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   556
  }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   557
  note ord = this
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   558
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   559
  show "p division_of (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   560
  proof (rule division_ofI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   561
    show "finite p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   562
      unfolding p_def by (auto intro!: finite_PiE)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   563
    {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   564
      fix k
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   565
      assume "k \<in> p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   566
      then obtain f where f: "f \<in> Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)}" and k: "k = ?B f"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   567
        by (auto simp: p_def)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   568
      then show "\<exists>a b. k = cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   569
        by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   570
      have "k \<subseteq> cbox a b \<and> k \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   571
      proof (simp add: k box_eq_empty subset_box not_less, safe)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   572
        fix i :: 'a
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   573
        assume i: "i \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   574
        with f have "f i = (a, c) \<or> f i = (c, d) \<or> f i = (d, b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   575
          by (auto simp: PiE_iff)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   576
        with i ord[of i]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   577
        show "a \<bullet> i \<le> fst (f i) \<bullet> i" "snd (f i) \<bullet> i \<le> b \<bullet> i" "fst (f i) \<bullet> i \<le> snd (f i) \<bullet> i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   578
          by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   579
      qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   580
      then show "k \<noteq> {}" "k \<subseteq> cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   581
        by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   582
      {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   583
        fix l
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   584
        assume "l \<in> p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   585
        then obtain g where g: "g \<in> Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)}" and l: "l = ?B g"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   586
          by (auto simp: p_def)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   587
        assume "l \<noteq> k"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   588
        have "\<exists>i\<in>Basis. f i \<noteq> g i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   589
        proof (rule ccontr)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   590
          assume "\<not> ?thesis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   591
          with f g have "f = g"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   592
            by (auto simp: PiE_iff extensional_def fun_eq_iff)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   593
          with \<open>l \<noteq> k\<close> show False
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   594
            by (simp add: l k)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   595
        qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   596
        then obtain i where *: "i \<in> Basis" "f i \<noteq> g i" ..
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   597
        then have "f i = (a, c) \<or> f i = (c, d) \<or> f i = (d, b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   598
                  "g i = (a, c) \<or> g i = (c, d) \<or> g i = (d, b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   599
          using f g by (auto simp: PiE_iff)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   600
        with * ord[of i] show "interior l \<inter> interior k = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   601
          by (auto simp add: l k interior_cbox disjoint_interval intro!: bexI[of _ i])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   602
      }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   603
      note \<open>k \<subseteq> cbox a b\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   604
    }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   605
    moreover
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   606
    {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   607
      fix x assume x: "x \<in> cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   608
      have "\<forall>i\<in>Basis. \<exists>l. x \<bullet> i \<in> {fst l \<bullet> i .. snd l \<bullet> i} \<and> l \<in> {(a, c), (c, d), (d, b)}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   609
      proof
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   610
        fix i :: 'a
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   611
        assume "i \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   612
        with x ord[of i]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   613
        have "(a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> c \<bullet> i) \<or> (c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i) \<or>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   614
            (d \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   615
          by (auto simp: cbox_def)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   616
        then show "\<exists>l. x \<bullet> i \<in> {fst l \<bullet> i .. snd l \<bullet> i} \<and> l \<in> {(a, c), (c, d), (d, b)}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   617
          by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   618
      qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   619
      then obtain f where
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   620
        f: "\<forall>i\<in>Basis. x \<bullet> i \<in> {fst (f i) \<bullet> i..snd (f i) \<bullet> i} \<and> f i \<in> {(a, c), (c, d), (d, b)}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   621
        unfolding bchoice_iff ..
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   622
      moreover from f have "restrict f Basis \<in> Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   623
        by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   624
      moreover from f have "x \<in> ?B (restrict f Basis)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   625
        by (auto simp: mem_box)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   626
      ultimately have "\<exists>k\<in>p. x \<in> k"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   627
        unfolding p_def by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   628
    }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   629
    ultimately show "\<Union>p = cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   630
      by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   631
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   632
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   633
66154
bc5e6461f759 Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents: 66113
diff changeset
   634
proposition partial_division_extend_interval:
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   635
  assumes "p division_of (\<Union>p)" "(\<Union>p) \<subseteq> cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   636
  obtains q where "p \<subseteq> q" "q division_of cbox a (b::'a::euclidean_space)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   637
proof (cases "p = {}")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   638
  case True
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   639
  obtain q where "q division_of (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   640
    by (rule elementary_interval)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   641
  then show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   642
    using True that by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   643
next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   644
  case False
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   645
  note p = division_ofD[OF assms(1)]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   646
  have div_cbox: "\<forall>k\<in>p. \<exists>q. q division_of cbox a b \<and> k \<in> q"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   647
  proof
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   648
    fix k
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   649
    assume kp: "k \<in> p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   650
    obtain c d where k: "k = cbox c d"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   651
      using p(4)[OF kp] by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   652
    have *: "cbox c d \<subseteq> cbox a b" "cbox c d \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   653
      using p(2,3)[OF kp, unfolded k] using assms(2)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   654
      by (blast intro: order.trans)+
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   655
    obtain q where "q division_of cbox a b" "cbox c d \<in> q"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   656
      by (rule partial_division_extend_1[OF *])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   657
    then show "\<exists>q. q division_of cbox a b \<and> k \<in> q"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   658
      unfolding k by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   659
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   660
  obtain q where q: "\<And>x. x \<in> p \<Longrightarrow> q x division_of cbox a b" "\<And>x. x \<in> p \<Longrightarrow> x \<in> q x"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   661
    using bchoice[OF div_cbox] by blast
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   662
  have "q x division_of \<Union>q x" if x: "x \<in> p" for x
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   663
    apply (rule division_ofI)
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   664
    using division_ofD[OF q(1)[OF x]] by auto
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   665
  then have di: "\<And>x. x \<in> p \<Longrightarrow> \<exists>d. d division_of \<Union>(q x - {x})"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   666
    by (meson Diff_subset division_of_subset)
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   667
  have "\<exists>d. d division_of \<Inter>((\<lambda>i. \<Union>(q i - {i})) ` p)"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   668
    apply (rule elementary_inters [OF finite_imageI[OF p(1)]])
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   669
    apply (auto dest: di simp: False elementary_inters [OF finite_imageI[OF p(1)]])
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   670
    done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   671
  then obtain d where d: "d division_of \<Inter>((\<lambda>i. \<Union>(q i - {i})) ` p)" ..
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   672
  have "d \<union> p division_of cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   673
  proof -
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   674
    have te: "\<And>S f T. S \<noteq> {} \<Longrightarrow> \<forall>i\<in>S. f i \<union> i = T \<Longrightarrow> T = \<Inter>(f ` S) \<union> \<Union>S" by auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   675
    have cbox_eq: "cbox a b = \<Inter>((\<lambda>i. \<Union>(q i - {i})) ` p) \<union> \<Union>p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   676
    proof (rule te[OF False], clarify)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   677
      fix i
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   678
      assume i: "i \<in> p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   679
      show "\<Union>(q i - {i}) \<union> i = cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   680
        using division_ofD(6)[OF q(1)[OF i]] using q(2)[OF i] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   681
    qed
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   682
    { fix K
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   683
      assume K: "K \<in> p"
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   684
      note qk=division_ofD[OF q(1)[OF K]]
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   685
      have *: "\<And>u T S. T \<inter> S = {} \<Longrightarrow> u \<subseteq> S \<Longrightarrow> u \<inter> T = {}"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   686
        by auto
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   687
      have "interior (\<Inter>i\<in>p. \<Union>(q i - {i})) \<inter> interior K = {}"
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   688
      proof (rule *[OF Int_interior_Union_intervals])
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   689
        show "\<And>T. T\<in>q K - {K} \<Longrightarrow> interior K \<inter> interior T = {}"
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   690
          using qk(5) using q(2)[OF K] by auto
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   691
        show "interior (\<Inter>i\<in>p. \<Union>(q i - {i})) \<subseteq> interior (\<Union>(q K - {K}))"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   692
          apply (rule interior_mono)+
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   693
          using K by auto
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   694
      qed (use qk in auto)} note [simp] = this
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   695
    show "d \<union> p division_of (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   696
      unfolding cbox_eq
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   697
      apply (rule division_disjoint_union[OF d assms(1)])
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   698
      apply (rule Int_interior_Union_intervals)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   699
      apply (rule p open_interior ballI)+
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   700
      apply simp_all
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   701
      done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   702
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   703
  then show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   704
    by (meson Un_upper2 that)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   705
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   706
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   707
lemma elementary_bounded[dest]:
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   708
  fixes S :: "'a::euclidean_space set"
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   709
  shows "p division_of S \<Longrightarrow> bounded S"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   710
  unfolding division_of_def by (metis bounded_Union bounded_cbox)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   711
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   712
lemma elementary_subset_cbox:
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   713
  "p division_of S \<Longrightarrow> \<exists>a b. S \<subseteq> cbox a (b::'a::euclidean_space)"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   714
  by (meson elementary_bounded bounded_subset_cbox)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   715
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   716
lemma division_union_intervals_exists:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   717
  fixes a b :: "'a::euclidean_space"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   718
  assumes "cbox a b \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   719
  obtains p where "(insert (cbox a b) p) division_of (cbox a b \<union> cbox c d)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   720
proof (cases "cbox c d = {}")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   721
  case True
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   722
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   723
    apply (rule that[of "{}"])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   724
    unfolding True
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   725
    using assms
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   726
    apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   727
    done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   728
next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   729
  case False
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   730
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   731
  proof (cases "cbox a b \<inter> cbox c d = {}")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   732
    case True
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   733
    then show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   734
      by (metis that False assms division_disjoint_union division_of_self insert_is_Un interior_Int interior_empty)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   735
  next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   736
    case False
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   737
    obtain u v where uv: "cbox a b \<inter> cbox c d = cbox u v"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   738
      unfolding Int_interval by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   739
    have uv_sub: "cbox u v \<subseteq> cbox c d" using uv by auto
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   740
    obtain p where pd: "p division_of cbox c d" and "cbox u v \<in> p"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   741
      by (rule partial_division_extend_1[OF uv_sub False[unfolded uv]])
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   742
    note p = this division_ofD[OF pd]
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   743
    have "interior (cbox a b \<inter> \<Union>(p - {cbox u v})) = interior(cbox u v \<inter> \<Union>(p - {cbox u v}))"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   744
      apply (rule arg_cong[of _ _ interior])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   745
      using p(8) uv by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   746
    also have "\<dots> = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   747
      unfolding interior_Int
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   748
      apply (rule Int_interior_Union_intervals)
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   749
      using p(6) p(7)[OF p(2)] \<open>finite p\<close>
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   750
      apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   751
      done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   752
    finally have [simp]: "interior (cbox a b) \<inter> interior (\<Union>(p - {cbox u v})) = {}" by simp
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   753
    have cbe: "cbox a b \<union> cbox c d = cbox a b \<union> \<Union>(p - {cbox u v})"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   754
      using p(8) unfolding uv[symmetric] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   755
    have "insert (cbox a b) (p - {cbox u v}) division_of cbox a b \<union> \<Union>(p - {cbox u v})"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   756
    proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   757
      have "{cbox a b} division_of cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   758
        by (simp add: assms division_of_self)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   759
      then show "insert (cbox a b) (p - {cbox u v}) division_of cbox a b \<union> \<Union>(p - {cbox u v})"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   760
        by (metis (no_types) Diff_subset \<open>interior (cbox a b) \<inter> interior (\<Union>(p - {cbox u v})) = {}\<close> division_disjoint_union division_of_subset insert_is_Un p(1) p(8))
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   761
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   762
    with that[of "p - {cbox u v}"] show ?thesis by (simp add: cbe)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   763
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   764
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   765
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   766
lemma division_of_Union:
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   767
  assumes "finite f"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   768
    and "\<And>p. p \<in> f \<Longrightarrow> p division_of (\<Union>p)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   769
    and "\<And>k1 k2. k1 \<in> \<Union>f \<Longrightarrow> k2 \<in> \<Union>f \<Longrightarrow> k1 \<noteq> k2 \<Longrightarrow> interior k1 \<inter> interior k2 = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   770
  shows "\<Union>f division_of \<Union>\<Union>f"
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   771
  using assms  by (auto intro!: division_ofI)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   772
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   773
lemma elementary_union_interval:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   774
  fixes a b :: "'a::euclidean_space"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   775
  assumes "p division_of \<Union>p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   776
  obtains q where "q division_of (cbox a b \<union> \<Union>p)"
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   777
proof (cases "p = {} \<or> cbox a b = {}")
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   778
  case True
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   779
  obtain p where "p division_of (cbox a b)"
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   780
    by (rule elementary_interval)
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   781
  then show thesis
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   782
    using True assms that by auto
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   783
next
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   784
  case False
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   785
  then have "p \<noteq> {}" "cbox a b \<noteq> {}"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   786
    by auto
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   787
  note pdiv = division_ofD[OF assms]
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   788
  show ?thesis
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   789
  proof (cases "interior (cbox a b) = {}")
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   790
    case True
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   791
    show ?thesis
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   792
      apply (rule that [of "insert (cbox a b) p", OF division_ofI])
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   793
      using pdiv(1-4) True False apply auto
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   794
       apply (metis IntI equals0D pdiv(5))
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   795
      by (metis disjoint_iff_not_equal pdiv(5))
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   796
  next
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   797
    case False  
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   798
    have "\<forall>K\<in>p. \<exists>q. (insert (cbox a b) q) division_of (cbox a b \<union> K)"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   799
    proof
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   800
      fix K
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   801
      assume kp: "K \<in> p"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   802
      from pdiv(4)[OF kp] obtain c d where "K = cbox c d" by blast
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   803
      then show "\<exists>q. (insert (cbox a b) q) division_of (cbox a b \<union> K)"
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   804
        by (meson \<open>cbox a b \<noteq> {}\<close> division_union_intervals_exists)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   805
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   806
    from bchoice[OF this] obtain q where "\<forall>x\<in>p. insert (cbox a b) (q x) division_of (cbox a b) \<union> x" ..
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   807
    note q = division_ofD[OF this[rule_format]]
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   808
    let ?D = "\<Union>{insert (cbox a b) (q K) | K. K \<in> p}"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   809
    show thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   810
    proof (rule that[OF division_ofI])
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   811
      have *: "{insert (cbox a b) (q K) |K. K \<in> p} = (\<lambda>K. insert (cbox a b) (q K)) ` p"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   812
        by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   813
      show "finite ?D"
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   814
        using "*" pdiv(1) q(1) by auto
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   815
      have lem1: "\<And>f s. \<Union>\<Union>(f ` s) = \<Union>((\<lambda>x. \<Union>(f x)) ` s)"
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   816
        by auto
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   817
      have lem2: "\<And>f s. f \<noteq> {} \<Longrightarrow> \<Union>{s \<union> t |t. t \<in> f} = s \<union> \<Union>f"
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   818
        by auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   819
      show "\<Union>?D = cbox a b \<union> \<Union>p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   820
        unfolding * lem1
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   821
        unfolding lem2[OF \<open>p \<noteq> {}\<close>, of "cbox a b", symmetric]
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   822
        using q(6)  by auto
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   823
      show "K \<subseteq> cbox a b \<union> \<Union>p" "K \<noteq> {}" if "K \<in> ?D" for K
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   824
        using q that by blast+
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   825
      show "\<exists>a b. K = cbox a b" if "K \<in> ?D" for K
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   826
        using q(4) that by auto
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   827
    next
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   828
      fix K' K
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   829
      assume K: "K \<in> ?D" and K': "K' \<in> ?D" "K \<noteq> K'"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   830
      obtain x where x: "K \<in> insert (cbox a b) (q x)" "x\<in>p"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   831
        using K by auto
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   832
      obtain x' where x': "K'\<in>insert (cbox a b) (q x')" "x'\<in>p"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   833
        using K' by auto
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   834
      show "interior K \<inter> interior K' = {}"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   835
      proof (cases "x = x' \<or> K  = cbox a b \<or> K' = cbox a b")
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   836
        case True
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   837
        then show ?thesis
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   838
          using True K' q(5) x' x by auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   839
      next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   840
        case False
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   841
        then have as': "K \<noteq> cbox a b" "K' \<noteq> cbox a b"
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   842
          by auto
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   843
        obtain c d where K: "K = cbox c d"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   844
          using q(4) x by blast
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   845
        have "interior K \<inter> interior (cbox a b) = {}"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   846
          using as' K'(2) q(5) x by blast
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   847
        then have "interior K \<subseteq> interior x"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   848
          by (metis \<open>interior (cbox a b) \<noteq> {}\<close> K q(2) x interior_subset_union_intervals)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   849
        moreover
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   850
        obtain c d where c_d: "K' = cbox c d"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   851
          using q(4)[OF x'(2,1)] by blast
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   852
        have "interior K' \<inter> interior (cbox a b) = {}"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   853
          using as'(2) q(5) x' by blast
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   854
        then have "interior K' \<subseteq> interior x'"
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   855
          by (metis \<open>interior (cbox a b) \<noteq> {}\<close> c_d interior_subset_union_intervals q(2) x'(1) x'(2))
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   856
        moreover have "interior x \<inter> interior x' = {}"
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   857
          by (meson False assms division_ofD(5) x'(2) x(2))
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   858
        ultimately show ?thesis
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   859
          using \<open>interior K \<subseteq> interior x\<close> \<open>interior K' \<subseteq> interior x'\<close> by auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   860
      qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   861
    qed
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   862
  qed
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   863
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   864
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   865
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   866
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   867
lemma elementary_unions_intervals:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   868
  assumes fin: "finite f"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   869
    and "\<And>s. s \<in> f \<Longrightarrow> \<exists>a b. s = cbox a (b::'a::euclidean_space)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   870
  obtains p where "p division_of (\<Union>f)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   871
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   872
  have "\<exists>p. p division_of (\<Union>f)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   873
  proof (induct_tac f rule:finite_subset_induct)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   874
    show "\<exists>p. p division_of \<Union>{}" using elementary_empty by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   875
  next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   876
    fix x F
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   877
    assume as: "finite F" "x \<notin> F" "\<exists>p. p division_of \<Union>F" "x\<in>f"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   878
    from this(3) obtain p where p: "p division_of \<Union>F" ..
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   879
    from assms(2)[OF as(4)] obtain a b where x: "x = cbox a b" by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   880
    have *: "\<Union>F = \<Union>p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   881
      using division_ofD[OF p] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   882
    show "\<exists>p. p division_of \<Union>insert x F"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   883
      using elementary_union_interval[OF p[unfolded *], of a b]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   884
      unfolding Union_insert x * by metis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   885
  qed (insert assms, auto)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   886
  then show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   887
    using that by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   888
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   889
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   890
lemma elementary_union:
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   891
  fixes S T :: "'a::euclidean_space set"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   892
  assumes "ps division_of S" "pt division_of T"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   893
  obtains p where "p division_of (S \<union> T)"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   894
proof -
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   895
  have *: "S \<union> T = \<Union>ps \<union> \<Union>pt"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   896
    using assms unfolding division_of_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   897
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   898
    apply (rule elementary_unions_intervals[of "ps \<union> pt"])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   899
    using assms apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   900
    by (simp add: * that)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   901
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   902
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   903
lemma partial_division_extend:
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   904
  fixes T :: "'a::euclidean_space set"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   905
  assumes "p division_of S"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   906
    and "q division_of T"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   907
    and "S \<subseteq> T"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   908
  obtains r where "p \<subseteq> r" and "r division_of T"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   909
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   910
  note divp = division_ofD[OF assms(1)] and divq = division_ofD[OF assms(2)]
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   911
  obtain a b where ab: "T \<subseteq> cbox a b"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   912
    using elementary_subset_cbox[OF assms(2)] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   913
  obtain r1 where "p \<subseteq> r1" "r1 division_of (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   914
    using assms
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   915
    by (metis ab dual_order.trans partial_division_extend_interval divp(6))
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   916
  note r1 = this division_ofD[OF this(2)]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   917
  obtain p' where "p' division_of \<Union>(r1 - p)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   918
    apply (rule elementary_unions_intervals[of "r1 - p"])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   919
    using r1(3,6)
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   920
      apply auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   921
    done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   922
  then obtain r2 where r2: "r2 division_of (\<Union>(r1 - p)) \<inter> (\<Union>q)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   923
    by (metis assms(2) divq(6) elementary_inter)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   924
  {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   925
    fix x
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   926
    assume x: "x \<in> T" "x \<notin> S"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   927
    then obtain R where r: "R \<in> r1" "x \<in> R"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   928
      unfolding r1 using ab
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   929
      by (meson division_contains r1(2) subsetCE)
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   930
    moreover have "R \<notin> p"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   931
    proof
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   932
      assume "R \<in> p"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   933
      then have "x \<in> S" using divp(2) r by auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   934
      then show False using x by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   935
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   936
    ultimately have "x\<in>\<Union>(r1 - p)" by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   937
  }
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   938
  then have Teq: "T = \<Union>p \<union> (\<Union>(r1 - p) \<inter> \<Union>q)"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   939
    unfolding divp divq using assms(3) by auto
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   940
  have "interior S \<inter> interior (\<Union>(r1-p)) = {}"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   941
  proof (rule Int_interior_Union_intervals)
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   942
    have *: "\<And>S. (\<And>x. x \<in> S \<Longrightarrow> False) \<Longrightarrow> S = {}"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   943
      by auto
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   944
    show "interior S \<inter> interior m = {}" if "m \<in> r1 - p" for m 
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   945
    proof -
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   946
      have "interior m \<inter> interior (\<Union>p) = {}"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   947
      proof (rule Int_interior_Union_intervals)
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   948
        show "\<And>T. T\<in>p \<Longrightarrow> interior m \<inter> interior T = {}"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   949
          by (metis DiffD1 DiffD2 that r1(1) r1(7) set_rev_mp)
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   950
      qed (use divp in auto)
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   951
      then show "interior S \<inter> interior m = {}"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   952
        unfolding divp by auto
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   953
    qed 
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   954
  qed (use r1 in auto)
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   955
  then have "interior S \<inter> interior (\<Union>(r1-p) \<inter> (\<Union>q)) = {}"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   956
    using interior_subset by auto
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   957
  then have div: "p \<union> r2 division_of \<Union>p \<union> \<Union>(r1 - p) \<inter> \<Union>q"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   958
    by (simp add: assms(1) division_disjoint_union divp(6) r2)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   959
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   960
    apply (rule that[of "p \<union> r2"])
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   961
     apply (auto simp: div Teq)
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   962
    done
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   963
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   964
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   965
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   966
lemma division_split:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   967
  fixes a :: "'a::euclidean_space"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   968
  assumes "p division_of (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   969
    and k: "k\<in>Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   970
  shows "{l \<inter> {x. x\<bullet>k \<le> c} | l. l \<in> p \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}} division_of(cbox a b \<inter> {x. x\<bullet>k \<le> c})"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   971
      (is "?p1 division_of ?I1")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   972
    and "{l \<inter> {x. x\<bullet>k \<ge> c} | l. l \<in> p \<and> l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}} division_of (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   973
      (is "?p2 division_of ?I2")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   974
proof (rule_tac[!] division_ofI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   975
  note p = division_ofD[OF assms(1)]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   976
  show "finite ?p1" "finite ?p2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   977
    using p(1) by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   978
  show "\<Union>?p1 = ?I1" "\<Union>?p2 = ?I2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   979
    unfolding p(6)[symmetric] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   980
  {
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   981
    fix K
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   982
    assume "K \<in> ?p1"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   983
    then obtain l where l: "K = l \<inter> {x. x \<bullet> k \<le> c}" "l \<in> p" "l \<inter> {x. x \<bullet> k \<le> c} \<noteq> {}"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   984
      by blast
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   985
    obtain u v where uv: "l = cbox u v"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   986
      using assms(1) l(2) by blast
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   987
    show "K \<subseteq> ?I1"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   988
      using l p(2) uv by force
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   989
    show  "K \<noteq> {}"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   990
      by (simp add: l)
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   991
    show  "\<exists>a b. K = cbox a b"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   992
      apply (simp add: l uv p(2-3)[OF l(2)])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   993
      apply (subst interval_split[OF k])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   994
      apply (auto intro: order.trans)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   995
      done
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   996
    fix K'
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   997
    assume "K' \<in> ?p1"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   998
    then obtain l' where l': "K' = l' \<inter> {x. x \<bullet> k \<le> c}" "l' \<in> p" "l' \<inter> {x. x \<bullet> k \<le> c} \<noteq> {}"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   999
      by blast
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1000
    assume "K \<noteq> K'"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1001
    then show "interior K \<inter> interior K' = {}"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1002
      unfolding l l' using p(5)[OF l(2) l'(2)] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1003
  }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1004
  {
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1005
    fix K
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1006
    assume "K \<in> ?p2"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1007
    then obtain l where l: "K = l \<inter> {x. c \<le> x \<bullet> k}" "l \<in> p" "l \<inter> {x. c \<le> x \<bullet> k} \<noteq> {}"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1008
      by blast
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1009
    obtain u v where uv: "l = cbox u v"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1010
      using l(2) p(4) by blast
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1011
    show "K \<subseteq> ?I2"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1012
      using l p(2) uv by force
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1013
    show  "K \<noteq> {}"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1014
      by (simp add: l)
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1015
    show  "\<exists>a b. K = cbox a b"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1016
      apply (simp add: l uv p(2-3)[OF l(2)])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1017
      apply (subst interval_split[OF k])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1018
      apply (auto intro: order.trans)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1019
      done
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1020
    fix K'
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1021
    assume "K' \<in> ?p2"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1022
    then obtain l' where l': "K' = l' \<inter> {x. c \<le> x \<bullet> k}" "l' \<in> p" "l' \<inter> {x. c \<le> x \<bullet> k} \<noteq> {}"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1023
      by blast
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1024
    assume "K \<noteq> K'"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1025
    then show "interior K \<inter> interior K' = {}"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1026
      unfolding l l' using p(5)[OF l(2) l'(2)] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1027
  }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1028
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1029
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1030
subsection \<open>Tagged (partial) divisions.\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1031
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1032
definition tagged_partial_division_of (infixr "tagged'_partial'_division'_of" 40)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1033
  where "s tagged_partial_division_of i \<longleftrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1034
    finite s \<and>
66113
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1035
    (\<forall>x K. (x, K) \<in> s \<longrightarrow> x \<in> K \<and> K \<subseteq> i \<and> (\<exists>a b. K = cbox a b)) \<and>
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1036
    (\<forall>x1 K1 x2 K2. (x1, K1) \<in> s \<and> (x2, K2) \<in> s \<and> (x1, K1) \<noteq> (x2, K2) \<longrightarrow>
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1037
      interior K1 \<inter> interior K2 = {})"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1038
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1039
lemma tagged_partial_division_ofD[dest]:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1040
  assumes "s tagged_partial_division_of i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1041
  shows "finite s"
66113
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1042
    and "\<And>x K. (x,K) \<in> s \<Longrightarrow> x \<in> K"
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1043
    and "\<And>x K. (x,K) \<in> s \<Longrightarrow> K \<subseteq> i"
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1044
    and "\<And>x K. (x,K) \<in> s \<Longrightarrow> \<exists>a b. K = cbox a b"
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1045
    and "\<And>x1 K1 x2 K2. (x1,K1) \<in> s \<Longrightarrow>
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1046
      (x2, K2) \<in> s \<Longrightarrow> (x1, K1) \<noteq> (x2, K2) \<Longrightarrow> interior K1 \<inter> interior K2 = {}"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1047
  using assms unfolding tagged_partial_division_of_def by blast+
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1048
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1049
definition tagged_division_of (infixr "tagged'_division'_of" 40)
66113
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1050
  where "s tagged_division_of i \<longleftrightarrow> s tagged_partial_division_of i \<and> (\<Union>{K. \<exists>x. (x,K) \<in> s} = i)"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1051
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1052
lemma tagged_division_of_finite: "s tagged_division_of i \<Longrightarrow> finite s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1053
  unfolding tagged_division_of_def tagged_partial_division_of_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1054
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1055
lemma tagged_division_of:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1056
  "s tagged_division_of i \<longleftrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1057
    finite s \<and>
66113
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1058
    (\<forall>x K. (x, K) \<in> s \<longrightarrow> x \<in> K \<and> K \<subseteq> i \<and> (\<exists>a b. K = cbox a b)) \<and>
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1059
    (\<forall>x1 K1 x2 K2. (x1, K1) \<in> s \<and> (x2, K2) \<in> s \<and> (x1, K1) \<noteq> (x2, K2) \<longrightarrow>
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1060
      interior K1 \<inter> interior K2 = {}) \<and>
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1061
    (\<Union>{K. \<exists>x. (x,K) \<in> s} = i)"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1062
  unfolding tagged_division_of_def tagged_partial_division_of_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1063
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1064
lemma tagged_division_ofI:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1065
  assumes "finite s"
66113
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1066
    and "\<And>x K. (x,K) \<in> s \<Longrightarrow> x \<in> K"
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1067
    and "\<And>x K. (x,K) \<in> s \<Longrightarrow> K \<subseteq> i"
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1068
    and "\<And>x K. (x,K) \<in> s \<Longrightarrow> \<exists>a b. K = cbox a b"
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1069
    and "\<And>x1 K1 x2 K2. (x1,K1) \<in> s \<Longrightarrow> (x2, K2) \<in> s \<Longrightarrow> (x1, K1) \<noteq> (x2, K2) \<Longrightarrow>
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1070
      interior K1 \<inter> interior K2 = {}"
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1071
    and "(\<Union>{K. \<exists>x. (x,K) \<in> s} = i)"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1072
  shows "s tagged_division_of i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1073
  unfolding tagged_division_of
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1074
  using assms  by fastforce
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1075
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1076
lemma tagged_division_ofD[dest]:  (*FIXME USE A LOCALE*)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1077
  assumes "s tagged_division_of i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1078
  shows "finite s"
66113
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1079
    and "\<And>x K. (x,K) \<in> s \<Longrightarrow> x \<in> K"
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1080
    and "\<And>x K. (x,K) \<in> s \<Longrightarrow> K \<subseteq> i"
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1081
    and "\<And>x K. (x,K) \<in> s \<Longrightarrow> \<exists>a b. K = cbox a b"
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1082
    and "\<And>x1 K1 x2 K2. (x1, K1) \<in> s \<Longrightarrow> (x2, K2) \<in> s \<Longrightarrow> (x1, K1) \<noteq> (x2, K2) \<Longrightarrow>
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1083
      interior K1 \<inter> interior K2 = {}"
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1084
    and "(\<Union>{K. \<exists>x. (x,K) \<in> s} = i)"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1085
  using assms unfolding tagged_division_of by blast+
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1086
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1087
lemma division_of_tagged_division:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1088
  assumes "s tagged_division_of i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1089
  shows "(snd ` s) division_of i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1090
proof (rule division_ofI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1091
  note assm = tagged_division_ofD[OF assms]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1092
  show "\<Union>(snd ` s) = i" "finite (snd ` s)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1093
    using assm by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1094
  fix k
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1095
  assume k: "k \<in> snd ` s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1096
  then obtain xk where xk: "(xk, k) \<in> s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1097
    by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1098
  then show "k \<subseteq> i" "k \<noteq> {}" "\<exists>a b. k = cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1099
    using assm by fastforce+
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1100
  fix k'
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1101
  assume k': "k' \<in> snd ` s" "k \<noteq> k'"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1102
  from this(1) obtain xk' where xk': "(xk', k') \<in> s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1103
    by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1104
  then show "interior k \<inter> interior k' = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1105
    using assm(5) k'(2) xk by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1106
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1107
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1108
lemma partial_division_of_tagged_division:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1109
  assumes "s tagged_partial_division_of i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1110
  shows "(snd ` s) division_of \<Union>(snd ` s)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1111
proof (rule division_ofI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1112
  note assm = tagged_partial_division_ofD[OF assms]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1113
  show "finite (snd ` s)" "\<Union>(snd ` s) = \<Union>(snd ` s)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1114
    using assm by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1115
  fix k
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1116
  assume k: "k \<in> snd ` s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1117
  then obtain xk where xk: "(xk, k) \<in> s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1118
    by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1119
  then show "k \<noteq> {}" "\<exists>a b. k = cbox a b" "k \<subseteq> \<Union>(snd ` s)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1120
    using assm by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1121
  fix k'
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1122
  assume k': "k' \<in> snd ` s" "k \<noteq> k'"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1123
  from this(1) obtain xk' where xk': "(xk', k') \<in> s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1124
    by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1125
  then show "interior k \<inter> interior k' = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1126
    using assm(5) k'(2) xk by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1127
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1128
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1129
lemma tagged_partial_division_subset:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1130
  assumes "s tagged_partial_division_of i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1131
    and "t \<subseteq> s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1132
  shows "t tagged_partial_division_of i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1133
  using assms
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1134
  unfolding tagged_partial_division_of_def
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1135
  using finite_subset[OF assms(2)]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1136
  by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1137
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1138
lemma tag_in_interval: "p tagged_division_of i \<Longrightarrow> (x, k) \<in> p \<Longrightarrow> x \<in> i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1139
  by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1140
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1141
lemma tagged_division_of_empty: "{} tagged_division_of {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1142
  unfolding tagged_division_of by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1143
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1144
lemma tagged_partial_division_of_trivial[simp]: "p tagged_partial_division_of {} \<longleftrightarrow> p = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1145
  unfolding tagged_partial_division_of_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1146
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1147
lemma tagged_division_of_trivial[simp]: "p tagged_division_of {} \<longleftrightarrow> p = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1148
  unfolding tagged_division_of by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1149
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1150
lemma tagged_division_of_self: "x \<in> cbox a b \<Longrightarrow> {(x,cbox a b)} tagged_division_of (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1151
  by (rule tagged_division_ofI) auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1152
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1153
lemma tagged_division_of_self_real: "x \<in> {a .. b::real} \<Longrightarrow> {(x,{a .. b})} tagged_division_of {a .. b}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1154
  unfolding box_real[symmetric]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1155
  by (rule tagged_division_of_self)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1156
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1157
lemma tagged_division_union:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1158
  assumes "p1 tagged_division_of s1"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1159
    and "p2 tagged_division_of s2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1160
    and "interior s1 \<inter> interior s2 = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1161
  shows "(p1 \<union> p2) tagged_division_of (s1 \<union> s2)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1162
proof (rule tagged_division_ofI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1163
  note p1 = tagged_division_ofD[OF assms(1)]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1164
  note p2 = tagged_division_ofD[OF assms(2)]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1165
  show "finite (p1 \<union> p2)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1166
    using p1(1) p2(1) by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1167
  show "\<Union>{k. \<exists>x. (x, k) \<in> p1 \<union> p2} = s1 \<union> s2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1168
    using p1(6) p2(6) by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1169
  fix x k
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1170
  assume xk: "(x, k) \<in> p1 \<union> p2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1171
  show "x \<in> k" "\<exists>a b. k = cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1172
    using xk p1(2,4) p2(2,4) by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1173
  show "k \<subseteq> s1 \<union> s2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1174
    using xk p1(3) p2(3) by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1175
  fix x' k'
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1176
  assume xk': "(x', k') \<in> p1 \<union> p2" "(x, k) \<noteq> (x', k')"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1177
  have *: "\<And>a b. a \<subseteq> s1 \<Longrightarrow> b \<subseteq> s2 \<Longrightarrow> interior a \<inter> interior b = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1178
    using assms(3) interior_mono by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1179
  show "interior k \<inter> interior k' = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1180
    apply (cases "(x, k) \<in> p1")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1181
    apply (meson "*" UnE assms(1) assms(2) p1(5) tagged_division_ofD(3) xk'(1) xk'(2))
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1182
    by (metis "*" UnE assms(1) assms(2) inf_sup_aci(1) p2(5) tagged_division_ofD(3) xk xk'(1) xk'(2))
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1183
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1184
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1185
lemma tagged_division_unions:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1186
  assumes "finite iset"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1187
    and "\<forall>i\<in>iset. pfn i tagged_division_of i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1188
    and "\<forall>i1\<in>iset. \<forall>i2\<in>iset. i1 \<noteq> i2 \<longrightarrow> interior(i1) \<inter> interior(i2) = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1189
  shows "\<Union>(pfn ` iset) tagged_division_of (\<Union>iset)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1190
proof (rule tagged_division_ofI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1191
  note assm = tagged_division_ofD[OF assms(2)[rule_format]]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1192
  show "finite (\<Union>(pfn ` iset))"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1193
    using assms by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1194
  have "\<Union>{k. \<exists>x. (x, k) \<in> \<Union>(pfn ` iset)} = \<Union>((\<lambda>i. \<Union>{k. \<exists>x. (x, k) \<in> pfn i}) ` iset)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1195
    by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1196
  also have "\<dots> = \<Union>iset"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1197
    using assm(6) by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1198
  finally show "\<Union>{k. \<exists>x. (x, k) \<in> \<Union>(pfn ` iset)} = \<Union>iset" .
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1199
  fix x k
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1200
  assume xk: "(x, k) \<in> \<Union>(pfn ` iset)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1201
  then obtain i where i: "i \<in> iset" "(x, k) \<in> pfn i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1202
    by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1203
  show "x \<in> k" "\<exists>a b. k = cbox a b" "k \<subseteq> \<Union>iset"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1204
    using assm(2-4)[OF i] using i(1) by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1205
  fix x' k'
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1206
  assume xk': "(x', k') \<in> \<Union>(pfn ` iset)" "(x, k) \<noteq> (x', k')"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1207
  then obtain i' where i': "i' \<in> iset" "(x', k') \<in> pfn i'"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1208
    by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1209
  have *: "\<And>a b. i \<noteq> i' \<Longrightarrow> a \<subseteq> i \<Longrightarrow> b \<subseteq> i' \<Longrightarrow> interior a \<inter> interior b = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1210
    using i(1) i'(1)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1211
    using assms(3)[rule_format] interior_mono
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1212
    by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1213
  show "interior k \<inter> interior k' = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1214
    apply (cases "i = i'")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1215
    using assm(5) i' i(2) xk'(2) apply blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1216
    using "*" assm(3) i' i by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1217
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1218
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1219
lemma tagged_partial_division_of_union_self:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1220
  assumes "p tagged_partial_division_of s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1221
  shows "p tagged_division_of (\<Union>(snd ` p))"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1222
  apply (rule tagged_division_ofI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1223
  using tagged_partial_division_ofD[OF assms]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1224
  apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1225
  done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1226
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1227
lemma tagged_division_of_union_self:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1228
  assumes "p tagged_division_of s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1229
  shows "p tagged_division_of (\<Union>(snd ` p))"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1230
  apply (rule tagged_division_ofI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1231
  using tagged_division_ofD[OF assms]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1232
  apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1233
  done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1234
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1235
subsection \<open>Functions closed on boxes: morphisms from boxes to monoids\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1236
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1237
text \<open>This auxiliary structure is used to sum up over the elements of a division. Main theorem is
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  1238
  \<open>operative_division\<close>. Instances for the monoid are @{typ "'a option"}, @{typ real}, and
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1239
  @{typ bool}.\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1240
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1241
paragraph \<open>Using additivity of lifted function to encode definedness.\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1242
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1243
definition lift_option :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a option \<Rightarrow> 'b option \<Rightarrow> 'c option"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1244
where
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1245
  "lift_option f a' b' = Option.bind a' (\<lambda>a. Option.bind b' (\<lambda>b. Some (f a b)))"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1246
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1247
lemma lift_option_simps[simp]:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1248
  "lift_option f (Some a) (Some b) = Some (f a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1249
  "lift_option f None b' = None"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1250
  "lift_option f a' None = None"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1251
  by (auto simp: lift_option_def)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1252
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1253
lemma comm_monoid_lift_option:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1254
  assumes "comm_monoid f z"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1255
  shows "comm_monoid (lift_option f) (Some z)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1256
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1257
  from assms interpret comm_monoid f z .
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1258
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1259
    by standard (auto simp: lift_option_def ac_simps split: bind_split)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1260
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1261
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1262
lemma comm_monoid_and: "comm_monoid HOL.conj True"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1263
  by standard auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1264
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1265
lemma comm_monoid_set_and: "comm_monoid_set HOL.conj True"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1266
  by (rule comm_monoid_set.intro) (fact comm_monoid_and)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1267
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1268
paragraph \<open>Operative\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1269
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1270
definition (in comm_monoid) operative :: "('b::euclidean_space set \<Rightarrow> 'a) \<Rightarrow> bool"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1271
  where "operative g \<longleftrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1272
    (\<forall>a b. box a b = {} \<longrightarrow> g (cbox a b) = \<^bold>1) \<and>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1273
    (\<forall>a b c. \<forall>k\<in>Basis. g (cbox a b) = g (cbox a b \<inter> {x. x\<bullet>k \<le> c}) \<^bold>* g (cbox a b \<inter> {x. x\<bullet>k \<ge> c}))"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1274
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1275
lemma (in comm_monoid) operativeD[dest]:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1276
  assumes "operative g"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1277
  shows "\<And>a b. box a b = {} \<Longrightarrow> g (cbox a b) = \<^bold>1"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1278
    and "\<And>a b c k. k \<in> Basis \<Longrightarrow> g (cbox a b) = g (cbox a b \<inter> {x. x\<bullet>k \<le> c}) \<^bold>* g (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1279
  using assms unfolding operative_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1280
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1281
lemma (in comm_monoid) operative_empty:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1282
  assumes g: "operative g" shows "g {} = \<^bold>1"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1283
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1284
  have *: "cbox One (-One) = ({}::'b set)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
  1285
    by (auto simp: box_eq_empty inner_sum_left inner_Basis sum.If_cases ex_in_conv)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1286
  moreover have "box One (-One) = ({}::'b set)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1287
    using box_subset_cbox[of One "-One"] by (auto simp: *)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1288
  ultimately show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1289
    using operativeD(1)[OF g, of One "-One"] by simp
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1290
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1291
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1292
definition "division_points (k::('a::euclidean_space) set) d =
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1293
   {(j,x). j \<in> Basis \<and> (interval_lowerbound k)\<bullet>j < x \<and> x < (interval_upperbound k)\<bullet>j \<and>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1294
     (\<exists>i\<in>d. (interval_lowerbound i)\<bullet>j = x \<or> (interval_upperbound i)\<bullet>j = x)}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1295
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1296
lemma division_points_finite:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1297
  fixes i :: "'a::euclidean_space set"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1298
  assumes "d division_of i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1299
  shows "finite (division_points i d)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1300
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1301
  note assm = division_ofD[OF assms]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1302
  let ?M = "\<lambda>j. {(j,x)|x. (interval_lowerbound i)\<bullet>j < x \<and> x < (interval_upperbound i)\<bullet>j \<and>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1303
    (\<exists>i\<in>d. (interval_lowerbound i)\<bullet>j = x \<or> (interval_upperbound i)\<bullet>j = x)}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1304
  have *: "division_points i d = \<Union>(?M ` Basis)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1305
    unfolding division_points_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1306
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1307
    unfolding * using assm by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1308
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1309
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1310
lemma division_points_subset:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1311
  fixes a :: "'a::euclidean_space"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1312
  assumes "d division_of (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1313
    and "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i"  "a\<bullet>k < c" "c < b\<bullet>k"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1314
    and k: "k \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1315
  shows "division_points (cbox a b \<inter> {x. x\<bullet>k \<le> c}) {l \<inter> {x. x\<bullet>k \<le> c} | l . l \<in> d \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}} \<subseteq>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1316
      division_points (cbox a b) d" (is ?t1)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1317
    and "division_points (cbox a b \<inter> {x. x\<bullet>k \<ge> c}) {l \<inter> {x. x\<bullet>k \<ge> c} | l . l \<in> d \<and> ~(l \<inter> {x. x\<bullet>k \<ge> c} = {})} \<subseteq>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1318
      division_points (cbox a b) d" (is ?t2)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1319
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1320
  note assm = division_ofD[OF assms(1)]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1321
  have *: "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1322
    "\<forall>i\<in>Basis. a\<bullet>i \<le> (\<Sum>i\<in>Basis. (if i = k then min (b \<bullet> k) c else  b \<bullet> i) *\<^sub>R i) \<bullet> i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1323
    "\<forall>i\<in>Basis. (\<Sum>i\<in>Basis. (if i = k then max (a \<bullet> k) c else a \<bullet> i) *\<^sub>R i) \<bullet> i \<le> b\<bullet>i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1324
    "min (b \<bullet> k) c = c" "max (a \<bullet> k) c = c"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1325
    using assms using less_imp_le by auto
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1326
   have "\<exists>i\<in>d. interval_lowerbound i \<bullet> x = y \<or> interval_upperbound i \<bullet> x = y"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1327
     if "a \<bullet> x < y" "y < (if x = k then c else b \<bullet> x)"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1328
        "interval_lowerbound i \<bullet> x = y \<or> interval_upperbound i \<bullet> x = y"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1329
        "i = l \<inter> {x. x \<bullet> k \<le> c}" "l \<in> d" "l \<inter> {x. x \<bullet> k \<le> c} \<noteq> {}"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1330
        "x \<in> Basis" for i l x y
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1331
  proof -
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1332
    obtain u v where l: "l = cbox u v"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1333
      using \<open>l \<in> d\<close> assms(1) by blast
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1334
    have *: "\<forall>i\<in>Basis. u \<bullet> i \<le> (\<Sum>i\<in>Basis. (if i = k then min (v \<bullet> k) c else v \<bullet> i) *\<^sub>R i) \<bullet> i"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1335
      using that(6) unfolding l interval_split[OF k] box_ne_empty that .
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1336
    have **: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1337
      using l using that(6) unfolding box_ne_empty[symmetric] by auto
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1338
    show ?thesis
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1339
      apply (rule bexI[OF _ \<open>l \<in> d\<close>])
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1340
      using that(1-3,5) \<open>x \<in> Basis\<close>
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1341
      unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] that
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1342
      apply (auto split: if_split_asm)
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1343
      done
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1344
  qed
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1345
  moreover have "\<And>x y. \<lbrakk>y < (if x = k then c else b \<bullet> x)\<rbrakk> \<Longrightarrow> y < b \<bullet> x"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1346
    using \<open>c < b\<bullet>k\<close> by (auto split: if_split_asm)
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1347
  ultimately show ?t1 
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1348
    unfolding division_points_def interval_split[OF k, of a b]
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1349
    unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)]
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1350
    unfolding * by force
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1351
  have "\<And>x y i l. (if x = k then c else a \<bullet> x) < y \<Longrightarrow> a \<bullet> x < y"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1352
    using \<open>a\<bullet>k < c\<close> by (auto split: if_split_asm)
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1353
  moreover have "\<exists>i\<in>d. interval_lowerbound i \<bullet> x = y \<or>
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1354
                       interval_upperbound i \<bullet> x = y"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1355
    if "(if x = k then c else a \<bullet> x) < y" "y < b \<bullet> x"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1356
      "interval_lowerbound i \<bullet> x = y \<or> interval_upperbound i \<bullet> x = y"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1357
      "i = l \<inter> {x. c \<le> x \<bullet> k}" "l \<in> d" "l \<inter> {x. c \<le> x \<bullet> k} \<noteq> {}"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1358
      "x \<in> Basis" for x y i l
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1359
  proof -
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1360
    obtain u v where l: "l = cbox u v"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1361
      using \<open>l \<in> d\<close> assm(4) by blast
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1362
    have *: "\<forall>i\<in>Basis. (\<Sum>i\<in>Basis. (if i = k then max (u \<bullet> k) c else u \<bullet> i) *\<^sub>R i) \<bullet> i \<le> v \<bullet> i"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1363
      using that(6) unfolding l interval_split[OF k] box_ne_empty that .
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1364
    have **: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1365
      using l using that(6) unfolding box_ne_empty[symmetric] by auto
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1366
    show "\<exists>i\<in>d. interval_lowerbound i \<bullet> x = y \<or> interval_upperbound i \<bullet> x = y"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1367
      apply (rule bexI[OF _ \<open>l \<in> d\<close>])
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1368
      using that(1-3,5) \<open>x \<in> Basis\<close>
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1369
      unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] that
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1370
      apply (auto split: if_split_asm)
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1371
      done
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1372
  qed
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1373
  ultimately show ?t2
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1374
    unfolding division_points_def interval_split[OF k, of a b]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1375
    unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1376
    unfolding *
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1377
    by force
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1378
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1379
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1380
lemma division_points_psubset:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1381
  fixes a :: "'a::euclidean_space"
66314
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  1382
  assumes d: "d division_of (cbox a b)"
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  1383
      and altb: "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i"  "a\<bullet>k < c" "c < b\<bullet>k"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1384
      and "l \<in> d"
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1385
      and disj: "interval_lowerbound l\<bullet>k = c \<or> interval_upperbound l\<bullet>k = c"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1386
      and k: "k \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1387
  shows "division_points (cbox a b \<inter> {x. x\<bullet>k \<le> c}) {l \<inter> {x. x\<bullet>k \<le> c} | l. l\<in>d \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}} \<subset>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1388
         division_points (cbox a b) d" (is "?D1 \<subset> ?D")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1389
    and "division_points (cbox a b \<inter> {x. x\<bullet>k \<ge> c}) {l \<inter> {x. x\<bullet>k \<ge> c} | l. l\<in>d \<and> l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}} \<subset>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1390
         division_points (cbox a b) d" (is "?D2 \<subset> ?D")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1391
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1392
  have ab: "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i"
66314
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  1393
    using altb by (auto intro!:less_imp_le)
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  1394
  obtain u v where l: "l = cbox u v"
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  1395
    using d \<open>l \<in> d\<close> by blast
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1396
  have uv: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i" "\<forall>i\<in>Basis. a\<bullet>i \<le> u\<bullet>i \<and> v\<bullet>i \<le> b\<bullet>i"
66314
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  1397
    apply (metis assms(5) box_ne_empty(1) cbox_division_memE d l)
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  1398
    by (metis assms(5) box_ne_empty(1) cbox_division_memE d l subset_box(1))
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1399
  have *: "interval_upperbound (cbox a b \<inter> {x. x \<bullet> k \<le> interval_upperbound l \<bullet> k}) \<bullet> k = interval_upperbound l \<bullet> k"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1400
          "interval_upperbound (cbox a b \<inter> {x. x \<bullet> k \<le> interval_lowerbound l \<bullet> k}) \<bullet> k = interval_lowerbound l \<bullet> k"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1401
    unfolding l interval_split[OF k] interval_bounds[OF uv(1)]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1402
    using uv[rule_format, of k] ab k
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1403
    by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1404
  have "\<exists>x. x \<in> ?D - ?D1"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1405
    using assms(3-)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1406
    unfolding division_points_def interval_bounds[OF ab]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1407
    apply -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1408
    apply (erule disjE)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1409
    apply (rule_tac x="(k,(interval_lowerbound l)\<bullet>k)" in exI, force simp add: *)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1410
    apply (rule_tac x="(k,(interval_upperbound l)\<bullet>k)" in exI, force simp add: *)
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1411
    done           
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1412
  moreover have "?D1 \<subseteq> ?D"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1413
    by (auto simp add: assms division_points_subset)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1414
  ultimately show "?D1 \<subset> ?D"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1415
    by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1416
  have *: "interval_lowerbound (cbox a b \<inter> {x. x \<bullet> k \<ge> interval_lowerbound l \<bullet> k}) \<bullet> k = interval_lowerbound l \<bullet> k"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1417
    "interval_lowerbound (cbox a b \<inter> {x. x \<bullet> k \<ge> interval_upperbound l \<bullet> k}) \<bullet> k = interval_upperbound l \<bullet> k"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1418
    unfolding l interval_split[OF k] interval_bounds[OF uv(1)]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1419
    using uv[rule_format, of k] ab k
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1420
    by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1421
  have "\<exists>x. x \<in> ?D - ?D2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1422
    using assms(3-)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1423
    unfolding division_points_def interval_bounds[OF ab]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1424
    apply -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1425
    apply (erule disjE)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1426
    apply (rule_tac x="(k,(interval_lowerbound l)\<bullet>k)" in exI, force simp add: *)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1427
    apply (rule_tac x="(k,(interval_upperbound l)\<bullet>k)" in exI, force simp add: *)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1428
    done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1429
  moreover have "?D2 \<subseteq> ?D"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1430
    by (auto simp add: assms division_points_subset)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1431
  ultimately show "?D2 \<subset> ?D"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1432
    by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1433
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1434
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1435
lemma division_split_left_inj:
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1436
  fixes S :: "'a::euclidean_space set"
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1437
  assumes div: "\<D> division_of S"
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1438
    and eq: "K1 \<inter> {x::'a. x\<bullet>k \<le> c} = K2 \<inter> {x. x\<bullet>k \<le> c}"
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1439
    and "K1 \<in> \<D>" "K2 \<in> \<D>" "K1 \<noteq> K2"
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1440
  shows "interior (K1 \<inter> {x. x\<bullet>k \<le> c}) = {}"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1441
proof -
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1442
  have "interior K2 \<inter> interior {a. a \<bullet> k \<le> c} = interior K1 \<inter> interior {a. a \<bullet> k \<le> c}"
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1443
    by (metis (no_types) eq interior_Int)
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1444
  moreover have "\<And>A. interior A \<inter> interior K2 = {} \<or> A = K2 \<or> A \<notin> \<D>"
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1445
    by (meson div \<open>K2 \<in> \<D>\<close> division_of_def)
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1446
  ultimately show ?thesis
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1447
    using \<open>K1 \<in> \<D>\<close> \<open>K1 \<noteq> K2\<close> by auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1448
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1449
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1450
lemma division_split_right_inj:
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1451
  fixes S :: "'a::euclidean_space set"
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1452
  assumes div: "\<D> division_of S"
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1453
    and eq: "K1 \<inter> {x::'a. x\<bullet>k \<ge> c} = K2 \<inter> {x. x\<bullet>k \<ge> c}"
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1454
    and "K1 \<in> \<D>" "K2 \<in> \<D>" "K1 \<noteq> K2"
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1455
  shows "interior (K1 \<inter> {x. x\<bullet>k \<ge> c}) = {}"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1456
proof -
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1457
  have "interior K2 \<inter> interior {a. a \<bullet> k \<ge> c} = interior K1 \<inter> interior {a. a \<bullet> k \<ge> c}"
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1458
    by (metis (no_types) eq interior_Int)
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1459
  moreover have "\<And>A. interior A \<inter> interior K2 = {} \<or> A = K2 \<or> A \<notin> \<D>"
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1460
    by (meson div \<open>K2 \<in> \<D>\<close> division_of_def)
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1461
  ultimately show ?thesis
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1462
    using \<open>K1 \<in> \<D>\<close> \<open>K1 \<noteq> K2\<close> by auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1463
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1464
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1465
lemma interval_doublesplit:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1466
  fixes a :: "'a::euclidean_space"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1467
  assumes "k \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1468
  shows "cbox a b \<inter> {x . \<bar>x\<bullet>k - c\<bar> \<le> (e::real)} =
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1469
    cbox (\<Sum>i\<in>Basis. (if i = k then max (a\<bullet>k) (c - e) else a\<bullet>i) *\<^sub>R i)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1470
     (\<Sum>i\<in>Basis. (if i = k then min (b\<bullet>k) (c + e) else b\<bullet>i) *\<^sub>R i)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1471
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1472
  have *: "\<And>x c e::real. \<bar>x - c\<bar> \<le> e \<longleftrightarrow> x \<ge> c - e \<and> x \<le> c + e"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1473
    by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1474
  have **: "\<And>s P Q. s \<inter> {x. P x \<and> Q x} = (s \<inter> {x. Q x}) \<inter> {x. P x}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1475
    by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1476
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1477
    unfolding * ** interval_split[OF assms] by (rule refl)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1478
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1479
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1480
lemma division_doublesplit:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1481
  fixes a :: "'a::euclidean_space"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1482
  assumes "p division_of (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1483
    and k: "k \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1484
  shows "(\<lambda>l. l \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> e}) ` {l\<in>p. l \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> e} \<noteq> {}}
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1485
         division_of  (cbox a b \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> e})"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1486
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1487
  have *: "\<And>x c. \<bar>x - c\<bar> \<le> e \<longleftrightarrow> x \<ge> c - e \<and> x \<le> c + e"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1488
    by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1489
  have **: "\<And>p q p' q'. p division_of q \<Longrightarrow> p = p' \<Longrightarrow> q = q' \<Longrightarrow> p' division_of q'"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1490
    by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1491
  note division_split(1)[OF assms, where c="c+e",unfolded interval_split[OF k]]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1492
  note division_split(2)[OF this, where c="c-e" and k=k,OF k]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1493
  then show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1494
    apply (rule **)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1495
    subgoal
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1496
      apply (simp add: abs_diff_le_iff field_simps Collect_conj_eq setcompr_eq_image[symmetric])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1497
      apply (rule equalityI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1498
      apply blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1499
      apply clarsimp
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1500
      apply (rule_tac x="l \<inter> {x. c + e \<ge> x \<bullet> k}" in exI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1501
      apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1502
      done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1503
    by (simp add: interval_split k interval_doublesplit)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1504
qed
66306
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1505
              
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1506
lemma (in comm_monoid_set) operative_division:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1507
  fixes g :: "'b::euclidean_space set \<Rightarrow> 'a"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1508
  assumes g: "operative g" and d: "d division_of (cbox a b)" shows "F g d = g (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1509
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1510
  define C where [abs_def]: "C = card (division_points (cbox a b) d)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1511
  then show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1512
    using d
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1513
  proof (induction C arbitrary: a b d rule: less_induct)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1514
    case (less a b d)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1515
    show ?case
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1516
    proof cases
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1517
      assume "box a b = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1518
      { fix k assume "k\<in>d"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1519
        then obtain a' b' where k: "k = cbox a' b'"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1520
          using division_ofD(4)[OF less.prems] by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1521
        with \<open>k\<in>d\<close> division_ofD(2)[OF less.prems] have "cbox a' b' \<subseteq> cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1522
          by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1523
        then have "box a' b' \<subseteq> box a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1524
          unfolding subset_box by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1525
        then have "g k = \<^bold>1"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1526
          using operativeD(1)[OF g, of a' b'] k by (simp add: \<open>box a b = {}\<close>) }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1527
      then show "box a b = {} \<Longrightarrow> F g d = g (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1528
        by (auto intro!: neutral simp: operativeD(1)[OF g])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1529
    next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1530
      assume "box a b \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1531
      then have ab: "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i" and ab': "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1532
        by (auto simp: box_ne_empty)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1533
      show "F g d = g (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1534
      proof (cases "division_points (cbox a b) d = {}")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1535
        case True
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1536
        { fix u v and j :: 'b
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1537
          assume j: "j \<in> Basis" and as: "cbox u v \<in> d"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1538
          then have "cbox u v \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1539
            using less.prems by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1540
          then have uv: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i" "u\<bullet>j \<le> v\<bullet>j"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1541
            using j unfolding box_ne_empty by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1542
          have *: "\<And>p r Q. \<not> j\<in>Basis \<or> p \<or> r \<or> (\<forall>x\<in>d. Q x) \<Longrightarrow> p \<or> r \<or> Q (cbox u v)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1543
            using as j by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1544
          have "(j, u\<bullet>j) \<notin> division_points (cbox a b) d"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1545
               "(j, v\<bullet>j) \<notin> division_points (cbox a b) d" using True by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1546
          note this[unfolded de_Morgan_conj division_points_def mem_Collect_eq split_conv interval_bounds[OF ab'] bex_simps]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1547
          note *[OF this(1)] *[OF this(2)] note this[unfolded interval_bounds[OF uv(1)]]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1548
          moreover
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1549
          have "a\<bullet>j \<le> u\<bullet>j" "v\<bullet>j \<le> b\<bullet>j"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1550
            using division_ofD(2,2,3)[OF \<open>d division_of cbox a b\<close> as]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1551
            apply (metis j subset_box(1) uv(1))
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1552
            by (metis \<open>cbox u v \<subseteq> cbox a b\<close> j subset_box(1) uv(1))
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1553
          ultimately have "u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = a\<bullet>j \<or> u\<bullet>j = b\<bullet>j \<and> v\<bullet>j = b\<bullet>j \<or> u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = b\<bullet>j"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1554
            unfolding not_less de_Morgan_disj using ab[rule_format,of j] uv(2) j by force }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1555
        then have d': "\<forall>i\<in>d. \<exists>u v. i = cbox u v \<and>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1556
          (\<forall>j\<in>Basis. u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = a\<bullet>j \<or> u\<bullet>j = b\<bullet>j \<and> v\<bullet>j = b\<bullet>j \<or> u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = b\<bullet>j)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1557
          unfolding forall_in_division[OF less.prems] by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1558
        have "(1/2) *\<^sub>R (a+b) \<in> cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1559
          unfolding mem_box using ab by (auto simp: inner_simps)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1560
        note this[unfolded division_ofD(6)[OF \<open>d division_of cbox a b\<close>,symmetric] Union_iff]
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1561
        then obtain i where i: "i \<in> d" "(1 / 2) *\<^sub>R (a + b) \<in> i" ..
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1562
        obtain u v where uv: "i = cbox u v" 
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1563
                     "\<forall>j\<in>Basis. u \<bullet> j = a \<bullet> j \<and> v \<bullet> j = a \<bullet> j \<or>
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1564
                                u \<bullet> j = b \<bullet> j \<and> v \<bullet> j = b \<bullet> j \<or>
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1565
                                u \<bullet> j = a \<bullet> j \<and> v \<bullet> j = b \<bullet> j"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1566
          using d' i(1) by auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1567
        have "cbox a b \<in> d"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1568
        proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1569
          have "u = a" "v = b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1570
            unfolding euclidean_eq_iff[where 'a='b]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1571
          proof safe
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1572
            fix j :: 'b
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1573
            assume j: "j \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1574
            note i(2)[unfolded uv mem_box,rule_format,of j]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1575
            then show "u \<bullet> j = a \<bullet> j" and "v \<bullet> j = b \<bullet> j"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1576
              using uv(2)[rule_format,of j] j by (auto simp: inner_simps)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1577
          qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1578
          then have "i = cbox a b" using uv by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1579
          then show ?thesis using i by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1580
        qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1581
        then have deq: "d = insert (cbox a b) (d - {cbox a b})"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1582
          by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1583
        have "F g (d - {cbox a b}) = \<^bold>1"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1584
        proof (intro neutral ballI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1585
          fix x
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1586
          assume x: "x \<in> d - {cbox a b}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1587
          then have "x\<in>d"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1588
            by auto note d'[rule_format,OF this]
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1589
          then obtain u v where uv: "x = cbox u v"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1590
                      "\<forall>j\<in>Basis. u \<bullet> j = a \<bullet> j \<and> v \<bullet> j = a \<bullet> j \<or>
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1591
                                 u \<bullet> j = b \<bullet> j \<and> v \<bullet> j = b \<bullet> j \<or>
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1592
                                 u \<bullet> j = a \<bullet> j \<and> v \<bullet> j = b \<bullet> j" 
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1593
            by blast
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1594
          have "u \<noteq> a \<or> v \<noteq> b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1595
            using x[unfolded uv] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1596
          then obtain j where "u\<bullet>j \<noteq> a\<bullet>j \<or> v\<bullet>j \<noteq> b\<bullet>j" and j: "j \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1597
            unfolding euclidean_eq_iff[where 'a='b] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1598
          then have "u\<bullet>j = v\<bullet>j"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1599
            using uv(2)[rule_format,OF j] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1600
          then have "box u v = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1601
            using j unfolding box_eq_empty by (auto intro!: bexI[of _ j])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1602
          then show "g x = \<^bold>1"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1603
            unfolding uv(1) by (rule operativeD(1)[OF g])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1604
        qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1605
        then show "F g d = g (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1606
          using division_ofD[OF less.prems]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1607
          apply (subst deq)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1608
          apply (subst insert)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1609
          apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1610
          done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1611
      next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1612
        case False
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1613
        then have "\<exists>x. x \<in> division_points (cbox a b) d"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1614
          by auto
66306
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1615
        then obtain k c 
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1616
          where "k \<in> Basis" "interval_lowerbound (cbox a b) \<bullet> k < c"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1617
                "c < interval_upperbound (cbox a b) \<bullet> k"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1618
                "\<exists>i\<in>d. interval_lowerbound i \<bullet> k = c \<or> interval_upperbound i \<bullet> k = c"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1619
          unfolding division_points_def by auto
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1620
        then obtain j where "a \<bullet> k < c" "c < b \<bullet> k" 
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1621
              and "j \<in> d" and j: "interval_lowerbound j \<bullet> k = c \<or> interval_upperbound j \<bullet> k = c"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1622
          by (metis division_of_trivial empty_iff interval_bounds' less.prems)
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1623
        let ?lec = "{x. x\<bullet>k \<le> c}" let ?gec = "{x. x\<bullet>k \<ge> c}"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1624
        define d1 where "d1 = {l \<inter> ?lec | l. l \<in> d \<and> l \<inter> ?lec \<noteq> {}}"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1625
        define d2 where "d2 = {l \<inter> ?gec | l. l \<in> d \<and> l \<inter> ?gec \<noteq> {}}"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1626
        define cb where "cb = (\<Sum>i\<in>Basis. (if i = k then c else b\<bullet>i) *\<^sub>R i)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1627
        define ca where "ca = (\<Sum>i\<in>Basis. (if i = k then c else a\<bullet>i) *\<^sub>R i)"
66306
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1628
        have "division_points (cbox a b \<inter> ?lec) {l \<inter> ?lec |l. l \<in> d \<and> l \<inter> ?lec \<noteq> {}} 
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1629
              \<subset> division_points (cbox a b) d"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1630
          by (rule division_points_psubset[OF \<open>d division_of cbox a b\<close> ab \<open>a \<bullet> k < c\<close> \<open>c < b \<bullet> k\<close> \<open>j \<in> d\<close> j \<open>k \<in> Basis\<close>])
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1631
        with division_points_finite[OF \<open>d division_of cbox a b\<close>] 
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1632
        have "card
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1633
         (division_points (cbox a b \<inter> ?lec) {l \<inter> ?lec |l. l \<in> d \<and> l \<inter> ?lec \<noteq> {}})
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1634
           < card (division_points (cbox a b) d)"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1635
          by (rule psubset_card_mono)
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1636
        moreover have "division_points (cbox a b \<inter> {x. c \<le> x \<bullet> k}) {l \<inter> {x. c \<le> x \<bullet> k} |l. l \<in> d \<and> l \<inter> {x. c \<le> x \<bullet> k} \<noteq> {}}
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1637
              \<subset> division_points (cbox a b) d"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1638
          by (rule division_points_psubset[OF \<open>d division_of cbox a b\<close> ab \<open>a \<bullet> k < c\<close> \<open>c < b \<bullet> k\<close> \<open>j \<in> d\<close> j \<open>k \<in> Basis\<close>])
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1639
        with division_points_finite[OF \<open>d division_of cbox a b\<close>] 
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1640
        have "card (division_points (cbox a b \<inter> ?gec) {l \<inter> ?gec |l. l \<in> d \<and> l \<inter> ?gec \<noteq> {}})
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1641
              < card (division_points (cbox a b) d)"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1642
          by (rule psubset_card_mono)
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1643
        ultimately have *: "F g d1 = g (cbox a b \<inter> ?lec)" "F g d2 = g (cbox a b \<inter> ?gec)"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1644
          unfolding interval_split[OF \<open>k \<in> Basis\<close>]
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1645
          apply (rule_tac[!] less.hyps)
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1646
          using division_split[OF \<open>d division_of cbox a b\<close>, where k=k and c=c] \<open>k \<in> Basis\<close>
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1647
          by (simp_all add: interval_split  d1_def d2_def division_points_finite[OF \<open>d division_of cbox a b\<close>])
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1648
        have fxk_le: "g (l \<inter> ?lec) = \<^bold>1" 
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1649
          if "l \<in> d" "y \<in> d" "l \<inter> ?lec = y \<inter> ?lec" "l \<noteq> y" for l y
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1650
        proof -
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1651
          obtain u v where leq: "l = cbox u v"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1652
            using \<open>l \<in> d\<close> less.prems by auto
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1653
          have "interior (cbox u v \<inter> ?lec) = {}"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1654
            using that division_split_left_inj leq less.prems by blast
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1655
          then show ?thesis
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1656
            unfolding leq interval_split[OF \<open>k \<in> Basis\<close>] using g by auto
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1657
        qed
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1658
        have fxk_ge: "g (l \<inter> {x. x \<bullet> k \<ge> c}) = \<^bold>1"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1659
          if "l \<in> d" "y \<in> d" "l \<inter> ?gec = y \<inter> ?gec" "l \<noteq> y" for l y
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1660
        proof -
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1661
          obtain u v where leq: "l = cbox u v"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1662
            using \<open>l \<in> d\<close> less.prems by auto
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1663
          have "interior (cbox u v \<inter> ?gec) = {}"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1664
            using that division_split_right_inj leq less.prems by blast
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1665
          then show ?thesis
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1666
            unfolding leq interval_split[OF \<open>k \<in> Basis\<close>] using g by auto
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1667
        qed
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1668
        have d1_alt: "d1 = (\<lambda>l. l \<inter> ?lec) ` {l \<in> d. l \<inter> ?lec \<noteq> {}}"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1669
          using d1_def by auto
66306
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1670
        have d2_alt: "d2 = (\<lambda>l. l \<inter> ?gec) ` {l \<in> d. l \<inter> ?gec \<noteq> {}}"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1671
          using d2_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1672
        have "g (cbox a b) = F g d1 \<^bold>* F g d2" (is "_ = ?prev")
66306
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1673
          unfolding * using g \<open>k \<in> Basis\<close> by blast
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1674
        also have "F g d1 = F (\<lambda>l. g (l \<inter> ?lec)) d"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1675
          unfolding d1_alt using division_of_finite[OF less.prems] fxk_le
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1676
          by (subst reindex_nontrivial) (auto intro!: mono_neutral_cong_left simp: operative_empty[OF g])
66306
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1677
        also have "F g d2 = F (\<lambda>l. g (l \<inter> ?gec)) d"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1678
          unfolding d2_alt using division_of_finite[OF less.prems] fxk_ge
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1679
          by (subst reindex_nontrivial) (auto intro!: mono_neutral_cong_left simp: operative_empty[OF g])
66306
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1680
        also have *: "\<forall>x\<in>d. g x = g (x \<inter> ?lec) \<^bold>* g (x \<inter> ?gec)"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1681
          unfolding forall_in_division[OF \<open>d division_of cbox a b\<close>]
66306
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1682
          using g \<open>k \<in> Basis\<close> by blast
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1683
        have "F (\<lambda>l. g (l \<inter> ?lec)) d \<^bold>* F (\<lambda>l. g (l \<inter> ?gec)) d = F g d"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1684
          using * by (simp add: distrib)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1685
        finally show ?thesis by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1686
      qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1687
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1688
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1689
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1690
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1691
lemma (in comm_monoid_set) over_tagged_division_lemma:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1692
  assumes "p tagged_division_of i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1693
    and "\<And>u v. cbox u v \<noteq> {} \<Longrightarrow> box u v = {} \<Longrightarrow> d (cbox u v) = \<^bold>1"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1694
  shows "F (\<lambda>(x,k). d k) p = F d (snd ` p)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1695
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1696
  have *: "(\<lambda>(x,k). d k) = d \<circ> snd"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1697
    unfolding o_def by (rule ext) auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1698
  note assm = tagged_division_ofD[OF assms(1)]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1699
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1700
    unfolding *
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1701
  proof (rule reindex_nontrivial[symmetric])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1702
    show "finite p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1703
      using assm by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1704
    fix x y
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1705
    assume "x\<in>p" "y\<in>p" "x\<noteq>y" "snd x = snd y"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1706
    obtain a b where ab: "snd x = cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1707
      using assm(4)[of "fst x" "snd x"] \<open>x\<in>p\<close> by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1708
    have "(fst x, snd y) \<in> p" "(fst x, snd y) \<noteq> y"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1709
      by (metis prod.collapse \<open>x\<in>p\<close> \<open>snd x = snd y\<close> \<open>x \<noteq> y\<close>)+
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1710
    with \<open>x\<in>p\<close> \<open>y\<in>p\<close> have "interior (snd x) \<inter> interior (snd y) = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1711
      by (intro assm(5)[of "fst x" _ "fst y"]) auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1712
    then have "box a b = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1713
      unfolding \<open>snd x = snd y\<close>[symmetric] ab by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1714
    then have "d (cbox a b) = \<^bold>1"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1715
      using assm(2)[of "fst x" "snd x"] \<open>x\<in>p\<close> ab[symmetric] by (intro assms(2)) auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1716
    then show "d (snd x) = \<^bold>1"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1717
      unfolding ab by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1718
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1719
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1720
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1721
lemma (in comm_monoid_set) operative_tagged_division:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1722
  assumes f: "operative g" and d: "d tagged_division_of (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1723
  shows "F (\<lambda>(x, l). g l) d = g (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1724
  unfolding d[THEN division_of_tagged_division, THEN operative_division[OF f], symmetric]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1725
  by (simp add: f[THEN operativeD(1)] over_tagged_division_lemma[OF d])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1726
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1727
lemma interval_real_split:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1728
  "{a .. b::real} \<inter> {x. x \<le> c} = {a .. min b c}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1729
  "{a .. b} \<inter> {x. c \<le> x} = {max a c .. b}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1730
  apply (metis Int_atLeastAtMostL1 atMost_def)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1731
  apply (metis Int_atLeastAtMostL2 atLeast_def)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1732
  done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1733
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1734
lemma (in comm_monoid) operative_1_lt:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1735
  "operative (g :: real set \<Rightarrow> 'a) \<longleftrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1736
    ((\<forall>a b. b \<le> a \<longrightarrow> g {a .. b} = \<^bold>1) \<and> (\<forall>a b c. a < c \<and> c < b \<longrightarrow> g {a .. c} \<^bold>* g {c .. b} = g {a .. b}))"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1737
  apply (simp add: operative_def atMost_def[symmetric] atLeast_def[symmetric])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1738
proof safe
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1739
  fix a b c :: real
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1740
  assume *: "\<forall>a b c. g {a..b} = g {a..min b c} \<^bold>* g {max a c..b}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1741
  assume "a < c" "c < b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1742
  with *[rule_format, of a b c] show "g {a..c} \<^bold>* g {c..b} = g {a..b}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1743
    by (simp add: less_imp_le min.absorb2 max.absorb2)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1744
next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1745
  fix a b c :: real
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1746
  assume as: "\<forall>a b. b \<le> a \<longrightarrow> g {a..b} = \<^bold>1"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1747
    "\<forall>a b c. a < c \<and> c < b \<longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1748
  from as(1)[rule_format, of 0 1] as(1)[rule_format, of a a for a] as(2)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1749
  have [simp]: "g {} = \<^bold>1" "\<And>a. g {a} = \<^bold>1"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1750
    "\<And>a b c. a < c \<Longrightarrow> c < b \<Longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1751
    by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1752
  show "g {a..b} = g {a..min b c} \<^bold>* g {max a c..b}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1753
    by (auto simp: min_def max_def le_less)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1754
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1755
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1756
lemma (in comm_monoid) operative_1_le:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1757
  "operative (g :: real set \<Rightarrow> 'a) \<longleftrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1758
    ((\<forall>a b. b \<le> a \<longrightarrow> g {a..b} = \<^bold>1) \<and> (\<forall>a b c. a \<le> c \<and> c \<le> b \<longrightarrow> g {a .. c} \<^bold>* g {c .. b} = g {a .. b}))"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1759
  unfolding operative_1_lt
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1760
proof safe
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1761
  fix a b c :: real
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1762
  assume as: "\<forall>a b c. a \<le> c \<and> c \<le> b \<longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}" "a < c" "c < b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1763
  show "g {a..c} \<^bold>* g {c..b} = g {a..b}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1764
    apply (rule as(1)[rule_format])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1765
    using as(2-)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1766
    apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1767
    done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1768
next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1769
  fix a b c :: real
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1770
  assume "\<forall>a b. b \<le> a \<longrightarrow> g {a .. b} = \<^bold>1"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1771
    and "\<forall>a b c. a < c \<and> c < b \<longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1772
    and "a \<le> c"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1773
    and "c \<le> b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1774
  note as = this[rule_format]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1775
  show "g {a..c} \<^bold>* g {c..b} = g {a..b}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1776
  proof (cases "c = a \<or> c = b")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1777
    case False
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1778
    then show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1779
      apply -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1780
      apply (subst as(2))
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1781
      using as(3-)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1782
      apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1783
      done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1784
  next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1785
    case True
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1786
    then show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1787
    proof
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1788
      assume *: "c = a"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1789
      then have "g {a .. c} = \<^bold>1"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1790
        apply -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1791
        apply (rule as(1)[rule_format])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1792
        apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1793
        done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1794
      then show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1795
        unfolding * by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1796
    next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1797
      assume *: "c = b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1798
      then have "g {c .. b} = \<^bold>1"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1799
        apply -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1800
        apply (rule as(1)[rule_format])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1801
        apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1802
        done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1803
      then show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1804
        unfolding * by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1805
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1806
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1807
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1808
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1809
lemma tagged_division_union_interval:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1810
  fixes a :: "'a::euclidean_space"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1811
  assumes "p1 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<le> (c::real)})"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1812
    and "p2 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1813
    and k: "k \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1814
  shows "(p1 \<union> p2) tagged_division_of (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1815
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1816
  have *: "cbox a b = (cbox a b \<inter> {x. x\<bullet>k \<le> c}) \<union> (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1817
    by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1818
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1819
    apply (subst *)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1820
    apply (rule tagged_division_union[OF assms(1-2)])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1821
    unfolding interval_split[OF k] interior_cbox
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1822
    using k
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1823
    apply (auto simp add: box_def elim!: ballE[where x=k])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1824
    done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1825
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1826
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1827
lemma tagged_division_union_interval_real:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1828
  fixes a :: real
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1829
  assumes "p1 tagged_division_of ({a .. b} \<inter> {x. x\<bullet>k \<le> (c::real)})"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1830
    and "p2 tagged_division_of ({a .. b} \<inter> {x. x\<bullet>k \<ge> c})"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1831
    and k: "k \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1832
  shows "(p1 \<union> p2) tagged_division_of {a .. b}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1833
  using assms
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1834
  unfolding box_real[symmetric]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1835
  by (rule tagged_division_union_interval)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1836
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1837
lemma tagged_division_split_left_inj:
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1838
  assumes d: "d tagged_division_of i"
66113
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1839
  and tags: "(x1, K1) \<in> d" "(x2, K2) \<in> d"
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1840
  and "K1 \<noteq> K2"
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1841
  and eq: "K1 \<inter> {x. x \<bullet> k \<le> c} = K2 \<inter> {x. x \<bullet> k \<le> c}"
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1842
    shows "interior (K1 \<inter> {x. x\<bullet>k \<le> c}) = {}"
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1843
proof -
66113
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1844
  have "interior (K1 \<inter> K2) = {} \<or> (x2, K2) = (x1, K1)"
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1845
    using tags d by (metis (no_types) interior_Int tagged_division_ofD(5))
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1846
  then show ?thesis
66113
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1847
    using eq \<open>K1 \<noteq> K2\<close> by (metis (no_types) inf_assoc inf_bot_left inf_left_idem interior_Int old.prod.inject)
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1848
qed
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1849
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1850
lemma tagged_division_split_right_inj:
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1851
  assumes d: "d tagged_division_of i"
66113
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1852
  and tags: "(x1, K1) \<in> d" "(x2, K2) \<in> d"
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1853
  and "K1 \<noteq> K2"
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1854
  and eq: "K1 \<inter> {x. x\<bullet>k \<ge> c} = K2 \<inter> {x. x\<bullet>k \<ge> c}"
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1855
    shows "interior (K1 \<inter> {x. x\<bullet>k \<ge> c}) = {}"
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1856
proof -
66113
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1857
  have "interior (K1 \<inter> K2) = {} \<or> (x2, K2) = (x1, K1)"
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1858
    using tags d by (metis (no_types) interior_Int tagged_division_ofD(5))
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1859
  then show ?thesis
66113
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1860
    using eq \<open>K1 \<noteq> K2\<close> by (metis (no_types) inf_assoc inf_bot_left inf_left_idem interior_Int old.prod.inject)
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1861
qed
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1862
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1863
subsection \<open>Special case of additivity we need for the FTC.\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1864
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1865
lemma additive_tagged_division_1:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1866
  fixes f :: "real \<Rightarrow> 'a::real_normed_vector"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1867
  assumes "a \<le> b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1868
    and "p tagged_division_of {a..b}"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
  1869
  shows "sum (\<lambda>(x,k). f(Sup k) - f(Inf k)) p = f b - f a"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1870
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1871
  let ?f = "(\<lambda>k::(real) set. if k = {} then 0 else f(interval_upperbound k) - f(interval_lowerbound k))"
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1872
  have p_td: "p tagged_division_of cbox a b"
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1873
    using assms(2) box_real(2) by presburger
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1874
  have *: "add.operative ?f"
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1875
    unfolding add.operative_1_lt box_eq_empty by auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1876
  have **: "cbox a b \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1877
    using assms(1) by auto
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1878
  then have "f b - f a = (\<Sum>(x, l)\<in>p. if l = {} then 0 else f (interval_upperbound l) - f (interval_lowerbound l))"
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1879
    proof -
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1880
      have "(if cbox a b = {} then 0 else f (interval_upperbound (cbox a b)) - f (interval_lowerbound (cbox a b))) = f b - f a"
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1881
        using assms by auto
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1882
      then show ?thesis
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1883
        using p_td assms by (simp add: "*" sum.operative_tagged_division)
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1884
    qed 
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1885
  then show ?thesis
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1886
    using assms by (auto intro!: sum.cong)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1887
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1888
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1889
lemma bgauge_existence_lemma: "(\<forall>x\<in>s. \<exists>d::real. 0 < d \<and> q d x) \<longleftrightarrow> (\<forall>x. \<exists>d>0. x\<in>s \<longrightarrow> q d x)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1890
  by (meson zero_less_one)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1891
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1892
lemma additive_tagged_division_1':
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1893
  fixes f :: "real \<Rightarrow> 'a::real_normed_vector"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1894
  assumes "a \<le> b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1895
    and "p tagged_division_of {a..b}"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
  1896
  shows "sum (\<lambda>(x,k). f (Sup k) - f(Inf k)) p = f b - f a"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1897
  using additive_tagged_division_1[OF _ assms(2), of f]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1898
  using assms(1)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1899
  by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1900
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1901
subsection \<open>Fine-ness of a partition w.r.t. a gauge.\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1902
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1903
definition fine  (infixr "fine" 46)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1904
  where "d fine s \<longleftrightarrow> (\<forall>(x,k) \<in> s. k \<subseteq> d x)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1905
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1906
lemma fineI:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1907
  assumes "\<And>x k. (x, k) \<in> s \<Longrightarrow> k \<subseteq> d x"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1908
  shows "d fine s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1909
  using assms unfolding fine_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1910
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1911
lemma fineD[dest]:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1912
  assumes "d fine s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1913
  shows "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> d x"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1914
  using assms unfolding fine_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1915
66192
e5b84854baa4 A few renamings and several tidied-up proofs
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
  1916
lemma fine_Int: "(\<lambda>x. d1 x \<inter> d2 x) fine p \<longleftrightarrow> d1 fine p \<and> d2 fine p"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1917
  unfolding fine_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1918
66192
e5b84854baa4 A few renamings and several tidied-up proofs
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
  1919
lemma fine_Inter:
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1920
 "(\<lambda>x. \<Inter>{f d x | d.  d \<in> s}) fine p \<longleftrightarrow> (\<forall>d\<in>s. (f d) fine p)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1921
  unfolding fine_def by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1922
66192
e5b84854baa4 A few renamings and several tidied-up proofs
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
  1923
lemma fine_Un: "d fine p1 \<Longrightarrow> d fine p2 \<Longrightarrow> d fine (p1 \<union> p2)"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1924
  unfolding fine_def by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1925
66192
e5b84854baa4 A few renamings and several tidied-up proofs
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
  1926
lemma fine_Union: "(\<And>p. p \<in> ps \<Longrightarrow> d fine p) \<Longrightarrow> d fine (\<Union>ps)"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1927
  unfolding fine_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1928
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1929
lemma fine_subset: "p \<subseteq> q \<Longrightarrow> d fine q \<Longrightarrow> d fine p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1930
  unfolding fine_def by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1931
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1932
subsection \<open>Some basic combining lemmas.\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1933
66192
e5b84854baa4 A few renamings and several tidied-up proofs
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
  1934
lemma tagged_division_Union_exists:
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1935
  assumes "finite iset"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1936
    and "\<forall>i\<in>iset. \<exists>p. p tagged_division_of i \<and> d fine p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1937
    and "\<forall>i1\<in>iset. \<forall>i2\<in>iset. i1 \<noteq> i2 \<longrightarrow> interior i1 \<inter> interior i2 = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1938
    and "\<Union>iset = i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1939
   obtains p where "p tagged_division_of i" and "d fine p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1940
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1941
  obtain pfn where pfn:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1942
    "\<And>x. x \<in> iset \<Longrightarrow> pfn x tagged_division_of x"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1943
    "\<And>x. x \<in> iset \<Longrightarrow> d fine pfn x"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1944
    using bchoice[OF assms(2)] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1945
  show thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1946
    apply (rule_tac p="\<Union>(pfn ` iset)" in that)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1947
    using assms(1) assms(3) assms(4) pfn(1) tagged_division_unions apply force
66192
e5b84854baa4 A few renamings and several tidied-up proofs
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
  1948
    by (metis (mono_tags, lifting) fine_Union imageE pfn(2))
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1949
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1950
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1951
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1952
subsection \<open>The set we're concerned with must be closed.\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1953
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1954
lemma division_of_closed:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1955
  fixes i :: "'n::euclidean_space set"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1956
  shows "s division_of i \<Longrightarrow> closed i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1957
  unfolding division_of_def by fastforce
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1958
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1959
subsection \<open>General bisection principle for intervals; might be useful elsewhere.\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1960
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1961
lemma interval_bisection_step:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1962
  fixes type :: "'a::euclidean_space"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1963
  assumes "P {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1964
    and "\<forall>s t. P s \<and> P t \<and> interior(s) \<inter> interior(t) = {} \<longrightarrow> P (s \<union> t)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1965
    and "\<not> P (cbox a (b::'a))"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1966
  obtains c d where "\<not> P (cbox c d)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1967
    and "\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> c\<bullet>i \<le> d\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i \<and> 2 * (d\<bullet>i - c\<bullet>i) \<le> b\<bullet>i - a\<bullet>i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1968
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1969
  have "cbox a b \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1970
    using assms(1,3) by metis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1971
  then have ab: "\<And>i. i\<in>Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1972
    by (force simp: mem_box)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1973
  { fix f
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1974
    have "\<lbrakk>finite f;
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1975
           \<And>s. s\<in>f \<Longrightarrow> P s;
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1976
           \<And>s. s\<in>f \<Longrightarrow> \<exists>a b. s = cbox a b;
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1977
           \<And>s t. s\<in>f \<Longrightarrow> t\<in>f \<Longrightarrow> s \<noteq> t \<Longrightarrow> interior s \<inter> interior t = {}\<rbrakk> \<Longrightarrow> P (\<Union>f)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1978
    proof (induct f rule: finite_induct)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1979
      case empty
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1980
      show ?case
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1981
        using assms(1) by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1982
    next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1983
      case (insert x f)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1984
      show ?case
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1985
        unfolding Union_insert
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1986
        apply (rule assms(2)[rule_format])
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
  1987
        using Int_interior_Union_intervals [of f "interior x"]
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1988
        apply (auto simp: insert)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1989
        by (metis IntI empty_iff insert.hyps(2) insert.prems(3) insert_iff)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1990
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1991
  } note UN_cases = this
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1992
  let ?A = "{cbox c d | c d::'a. \<forall>i\<in>Basis. (c\<bullet>i = a\<bullet>i) \<and> (d\<bullet>i = (a\<bullet>i + b\<bullet>i) / 2) \<or>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1993
    (c\<bullet>i = (a\<bullet>i + b\<bullet>i) / 2) \<and> (d\<bullet>i = b\<bullet>i)}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1994
  let ?PP = "\<lambda>c d. \<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> c\<bullet>i \<le> d\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i \<and> 2 * (d\<bullet>i - c\<bullet>i) \<le> b\<bullet>i - a\<bullet>i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1995
  {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1996
    presume "\<forall>c d. ?PP c d \<longrightarrow> P (cbox c d) \<Longrightarrow> False"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1997
    then show thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1998
      unfolding atomize_not not_all
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1999
      by (blast intro: that)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2000
  }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2001
  assume as: "\<forall>c d. ?PP c d \<longrightarrow> P (cbox c d)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2002
  have "P (\<Union>?A)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2003
  proof (rule UN_cases)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2004
    let ?B = "(\<lambda>s. cbox (\<Sum>i\<in>Basis. (if i \<in> s then a\<bullet>i else (a\<bullet>i + b\<bullet>i) / 2) *\<^sub>R i::'a)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2005
      (\<Sum>i\<in>Basis. (if i \<in> s then (a\<bullet>i + b\<bullet>i) / 2 else b\<bullet>i) *\<^sub>R i)) ` {s. s \<subseteq> Basis}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2006
    have "?A \<subseteq> ?B"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2007
    proof
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2008
      fix x
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2009
      assume "x \<in> ?A"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2010
      then obtain c d
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2011
        where x:  "x = cbox c d"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2012
                  "\<And>i. i \<in> Basis \<Longrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2013
                        c \<bullet> i = a \<bullet> i \<and> d \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<or>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2014
                        c \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<and> d \<bullet> i = b \<bullet> i" by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2015
      show "x \<in> ?B"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2016
        unfolding image_iff x
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2017
        apply (rule_tac x="{i. i\<in>Basis \<and> c\<bullet>i = a\<bullet>i}" in bexI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2018
        apply (rule arg_cong2 [where f = cbox])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2019
        using x(2) ab
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2020
        apply (auto simp add: euclidean_eq_iff[where 'a='a])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2021
        by fastforce
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2022
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2023
    then show "finite ?A"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2024
      by (rule finite_subset) auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2025
  next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2026
    fix s
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2027
    assume "s \<in> ?A"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2028
    then obtain c d
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2029
      where s: "s = cbox c d"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2030
               "\<And>i. i \<in> Basis \<Longrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2031
                     c \<bullet> i = a \<bullet> i \<and> d \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<or>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2032
                     c \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<and> d \<bullet> i = b \<bullet> i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2033
      by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2034
    show "P s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2035
      unfolding s
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2036
      apply (rule as[rule_format])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2037
      using ab s(2) by force
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2038
    show "\<exists>a b. s = cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2039
      unfolding s by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2040
    fix t
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2041
    assume "t \<in> ?A"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2042
    then obtain e f where t:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2043
      "t = cbox e f"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2044
      "\<And>i. i \<in> Basis \<Longrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2045
        e \<bullet> i = a \<bullet> i \<and> f \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<or>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2046
        e \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<and> f \<bullet> i = b \<bullet> i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2047
      by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2048
    assume "s \<noteq> t"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2049
    then have "\<not> (c = e \<and> d = f)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2050
      unfolding s t by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2051
    then obtain i where "c\<bullet>i \<noteq> e\<bullet>i \<or> d\<bullet>i \<noteq> f\<bullet>i" and i': "i \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2052
      unfolding euclidean_eq_iff[where 'a='a] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2053
    then have i: "c\<bullet>i \<noteq> e\<bullet>i" "d\<bullet>i \<noteq> f\<bullet>i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2054
      using s(2) t(2) apply fastforce
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2055
      using t(2)[OF i'] \<open>c \<bullet> i \<noteq> e \<bullet> i \<or> d \<bullet> i \<noteq> f \<bullet> i\<close> i' s(2) t(2) by fastforce
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2056
    have *: "\<And>s t. (\<And>a. a \<in> s \<Longrightarrow> a \<in> t \<Longrightarrow> False) \<Longrightarrow> s \<inter> t = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2057
      by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2058
    show "interior s \<inter> interior t = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2059
      unfolding s t interior_cbox
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2060
    proof (rule *)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2061
      fix x
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2062
      assume "x \<in> box c d" "x \<in> box e f"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2063
      then have x: "c\<bullet>i < d\<bullet>i" "e\<bullet>i < f\<bullet>i" "c\<bullet>i < f\<bullet>i" "e\<bullet>i < d\<bullet>i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2064
        unfolding mem_box using i'
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2065
        by force+
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2066
      show False  using s(2)[OF i']
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2067
      proof safe
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2068
        assume as: "c \<bullet> i = a \<bullet> i" "d \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2069
        show False
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2070
          using t(2)[OF i'] and i x unfolding as by (fastforce simp add:field_simps)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2071
      next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2072
        assume as: "c \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2" "d \<bullet> i = b \<bullet> i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2073
        show False
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2074
          using t(2)[OF i'] and i x unfolding as by(fastforce simp add:field_simps)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2075
      qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2076
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2077
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2078
  also have "\<Union>?A = cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2079
  proof (rule set_eqI,rule)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2080
    fix x
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2081
    assume "x \<in> \<Union>?A"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2082
    then obtain c d where x:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2083
      "x \<in> cbox c d"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2084
      "\<And>i. i \<in> Basis \<Longrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2085
        c \<bullet> i = a \<bullet> i \<and> d \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<or>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2086
        c \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<and> d \<bullet> i = b \<bullet> i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2087
      by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2088
    show "x\<in>cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2089
      unfolding mem_box
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2090
    proof safe
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2091
      fix i :: 'a
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2092
      assume i: "i \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2093
      then show "a \<bullet> i \<le> x \<bullet> i" "x \<bullet> i \<le> b \<bullet> i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2094
        using x(2)[OF i] x(1)[unfolded mem_box,THEN bspec, OF i] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2095
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2096
  next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2097
    fix x
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2098
    assume x: "x \<in> cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2099
    have "\<forall>i\<in>Basis.
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2100
      \<exists>c d. (c = a\<bullet>i \<and> d = (a\<bullet>i + b\<bullet>i) / 2 \<or> c = (a\<bullet>i + b\<bullet>i) / 2 \<and> d = b\<bullet>i) \<and> c\<le>x\<bullet>i \<and> x\<bullet>i \<le> d"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2101
      (is "\<forall>i\<in>Basis. \<exists>c d. ?P i c d")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2102
      unfolding mem_box
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2103
    proof
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2104
      fix i :: 'a
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2105
      assume i: "i \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2106
      have "?P i (a\<bullet>i) ((a \<bullet> i + b \<bullet> i) / 2) \<or> ?P i ((a \<bullet> i + b \<bullet> i) / 2) (b\<bullet>i)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2107
        using x[unfolded mem_box,THEN bspec, OF i] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2108
      then show "\<exists>c d. ?P i c d"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2109
        by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2110
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2111
    then show "x\<in>\<Union>?A"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2112
      unfolding Union_iff Bex_def mem_Collect_eq choice_Basis_iff
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2113
      apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2114
      apply (rule_tac x="cbox xa xaa" in exI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2115
      unfolding mem_box
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2116
      apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2117
      done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2118
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2119
  finally show False
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2120
    using assms by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2121
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2122
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2123
lemma interval_bisection:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2124
  fixes type :: "'a::euclidean_space"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2125
  assumes "P {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2126
    and "(\<forall>s t. P s \<and> P t \<and> interior(s) \<inter> interior(t) = {} \<longrightarrow> P(s \<union> t))"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2127
    and "\<not> P (cbox a (b::'a))"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2128
  obtains x where "x \<in> cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2129
    and "\<forall>e>0. \<exists>c d. x \<in> cbox c d \<and> cbox c d \<subseteq> ball x e \<and> cbox c d \<subseteq> cbox a b \<and> \<not> P (cbox c d)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2130
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2131
  have "\<forall>x. \<exists>y. \<not> P (cbox (fst x) (snd x)) \<longrightarrow> (\<not> P (cbox (fst y) (snd y)) \<and>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2132
    (\<forall>i\<in>Basis. fst x\<bullet>i \<le> fst y\<bullet>i \<and> fst y\<bullet>i \<le> snd y\<bullet>i \<and> snd y\<bullet>i \<le> snd x\<bullet>i \<and>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2133
       2 * (snd y\<bullet>i - fst y\<bullet>i) \<le> snd x\<bullet>i - fst x\<bullet>i))" (is "\<forall>x. ?P x")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2134
  proof
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2135
    show "?P x" for x
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2136
    proof (cases "P (cbox (fst x) (snd x))")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2137
      case True
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2138
      then show ?thesis by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2139
    next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2140
      case as: False
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2141
      obtain c d where "\<not> P (cbox c d)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2142
        "\<forall>i\<in>Basis.
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2143
           fst x \<bullet> i \<le> c \<bullet> i \<and>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2144
           c \<bullet> i \<le> d \<bullet> i \<and>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2145
           d \<bullet> i \<le> snd x \<bullet> i \<and>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2146
           2 * (d \<bullet> i - c \<bullet> i) \<le> snd x \<bullet> i - fst x \<bullet> i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2147
        by (rule interval_bisection_step[of P, OF assms(1-2) as])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2148
      then show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2149
        apply -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2150
        apply (rule_tac x="(c,d)" in exI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2151
        apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2152
        done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2153
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2154
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2155
  then obtain f where f:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2156
    "\<forall>x.
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2157
      \<not> P (cbox (fst x) (snd x)) \<longrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2158
      \<not> P (cbox (fst (f x)) (snd (f x))) \<and>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2159
        (\<forall>i\<in>Basis.
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2160
            fst x \<bullet> i \<le> fst (f x) \<bullet> i \<and>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2161
            fst (f x) \<bullet> i \<le> snd (f x) \<bullet> i \<and>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2162
            snd (f x) \<bullet> i \<le> snd x \<bullet> i \<and>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2163
            2 * (snd (f x) \<bullet> i - fst (f x) \<bullet> i) \<le> snd x \<bullet> i - fst x \<bullet> i)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2164
    apply -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2165
    apply (drule choice)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2166
    apply blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2167
    done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2168
  define AB A B where ab_def: "AB n = (f ^^ n) (a,b)" "A n = fst(AB n)" "B n = snd(AB n)" for n
66193
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2169
  have [simp]: "A 0 = a" "B 0 = b" and ABRAW: "\<And>n. \<not> P (cbox (A(Suc n)) (B(Suc n))) \<and>
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2170
    (\<forall>i\<in>Basis. A(n)\<bullet>i \<le> A(Suc n)\<bullet>i \<and> A(Suc n)\<bullet>i \<le> B(Suc n)\<bullet>i \<and> B(Suc n)\<bullet>i \<le> B(n)\<bullet>i \<and>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2171
    2 * (B(Suc n)\<bullet>i - A(Suc n)\<bullet>i) \<le> B(n)\<bullet>i - A(n)\<bullet>i)" (is "\<And>n. ?P n")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2172
  proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2173
    show "A 0 = a" "B 0 = b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2174
      unfolding ab_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2175
    note S = ab_def funpow.simps o_def id_apply
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2176
    show "?P n" for n
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2177
    proof (induct n)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2178
      case 0
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2179
      then show ?case
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2180
        unfolding S
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2181
        apply (rule f[rule_format]) using assms(3)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2182
        apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2183
        done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2184
    next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2185
      case (Suc n)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2186
      show ?case
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2187
        unfolding S
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2188
        apply (rule f[rule_format])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2189
        using Suc
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2190
        unfolding S
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2191
        apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2192
        done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2193
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2194
  qed
66193
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2195
  then have AB: "A(n)\<bullet>i \<le> A(Suc n)\<bullet>i" "A(Suc n)\<bullet>i \<le> B(Suc n)\<bullet>i" 
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2196
                 "B(Suc n)\<bullet>i \<le> B(n)\<bullet>i" "2 * (B(Suc n)\<bullet>i - A(Suc n)\<bullet>i) \<le> B(n)\<bullet>i - A(n)\<bullet>i" 
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2197
                if "i\<in>Basis" for i n
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2198
    using that by blast+
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2199
  have notPAB: "\<not> P (cbox (A(Suc n)) (B(Suc n)))" for n
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2200
    using ABRAW by blast
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2201
  have interv: "\<exists>n. \<forall>x\<in>cbox (A n) (B n). \<forall>y\<in>cbox (A n) (B n). dist x y < e"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2202
    if e: "0 < e" for e
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2203
  proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2204
    obtain n where n: "(\<Sum>i\<in>Basis. b \<bullet> i - a \<bullet> i) / e < 2 ^ n"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
  2205
      using real_arch_pow[of 2 "(sum (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis) / e"] by auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2206
    show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2207
    proof (rule exI [where x=n], clarify)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2208
      fix x y
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2209
      assume xy: "x\<in>cbox (A n) (B n)" "y\<in>cbox (A n) (B n)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
  2210
      have "dist x y \<le> sum (\<lambda>i. \<bar>(x - y)\<bullet>i\<bar>) Basis"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2211
        unfolding dist_norm by(rule norm_le_l1)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
  2212
      also have "\<dots> \<le> sum (\<lambda>i. B n\<bullet>i - A n\<bullet>i) Basis"
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
  2213
      proof (rule sum_mono)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2214
        fix i :: 'a
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2215
        assume i: "i \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2216
        show "\<bar>(x - y) \<bullet> i\<bar> \<le> B n \<bullet> i - A n \<bullet> i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2217
          using xy[unfolded mem_box,THEN bspec, OF i]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2218
          by (auto simp: inner_diff_left)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2219
      qed
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
  2220
      also have "\<dots> \<le> sum (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis / 2^n"
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
  2221
        unfolding sum_divide_distrib
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
  2222
      proof (rule sum_mono)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2223
        show "B n \<bullet> i - A n \<bullet> i \<le> (b \<bullet> i - a \<bullet> i) / 2 ^ n" if i: "i \<in> Basis" for i
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2224
        proof (induct n)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2225
          case 0
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2226
          then show ?case
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2227
            unfolding AB by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2228
        next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2229
          case (Suc n)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2230
          have "B (Suc n) \<bullet> i - A (Suc n) \<bullet> i \<le> (B n \<bullet> i - A n \<bullet> i) / 2"
66193
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2231
            using AB(3) that
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2232
            using AB(4)[of i n] using i by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2233
          also have "\<dots> \<le> (b \<bullet> i - a \<bullet> i) / 2 ^ Suc n"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2234
            using Suc by (auto simp add: field_simps)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2235
          finally show ?case .
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2236
        qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2237
      qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2238
      also have "\<dots> < e"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2239
        using n using e by (auto simp add: field_simps)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2240
      finally show "dist x y < e" .
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2241
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2242
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2243
  {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2244
    fix n m :: nat
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2245
    assume "m \<le> n" then have "cbox (A n) (B n) \<subseteq> cbox (A m) (B m)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2246
    proof (induction rule: inc_induct)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2247
      case (step i)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2248
      show ?case
66193
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2249
        using AB by (intro order_trans[OF step.IH] subset_box_imp) auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2250
    qed simp
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2251
  } note ABsubset = this
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2252
  have "\<exists>a. \<forall>n. a\<in> cbox (A n) (B n)"
66193
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2253
  proof (rule decreasing_closed_nest)
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2254
    show "\<forall>n. closed (cbox (A n) (B n))"
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2255
      by (simp add: closed_cbox)
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2256
    show "\<forall>n. cbox (A n) (B n) \<noteq> {}"
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2257
      by (meson AB dual_order.trans interval_not_empty)
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2258
  qed (use ABsubset interv in auto)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2259
  then obtain x0 where x0: "\<And>n. x0 \<in> cbox (A n) (B n)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2260
    by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2261
  show thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2262
  proof (rule that[rule_format, of x0])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2263
    show "x0\<in>cbox a b"
66193
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2264
      using \<open>A 0 = a\<close> \<open>B 0 = b\<close> x0 by blast
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2265
    fix e :: real
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2266
    assume "e > 0"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2267
    from interv[OF this] obtain n
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2268
      where n: "\<forall>x\<in>cbox (A n) (B n). \<forall>y\<in>cbox (A n) (B n). dist x y < e" ..
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2269
    have "\<not> P (cbox (A n) (B n))"
66193
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2270
    proof (cases "0 < n")
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2271
      case True then show ?thesis 
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2272
        by (metis Suc_pred' notPAB) 
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2273
    next
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2274
      case False then show ?thesis
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2275
        using \<open>A 0 = a\<close> \<open>B 0 = b\<close> \<open>\<not> P (cbox a b)\<close> by blast
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2276
    qed
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2277
    moreover have "cbox (A n) (B n) \<subseteq> ball x0 e"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2278
      using n using x0[of n] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2279
    moreover have "cbox (A n) (B n) \<subseteq> cbox a b"
66193
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2280
      using ABsubset \<open>A 0 = a\<close> \<open>B 0 = b\<close> by blast
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2281
    ultimately show "\<exists>c d. x0 \<in> cbox c d \<and> cbox c d \<subseteq> ball x0 e \<and> cbox c d \<subseteq> cbox a b \<and> \<not> P (cbox c d)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2282
      apply (rule_tac x="A n" in exI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2283
      apply (rule_tac x="B n" in exI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2284
      apply (auto simp: x0)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2285
      done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2286
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2287
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2288
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2289
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2290
subsection \<open>Cousin's lemma.\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2291
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2292
lemma fine_division_exists:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2293
  fixes a b :: "'a::euclidean_space"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2294
  assumes "gauge g"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2295
  obtains p where "p tagged_division_of (cbox a b)" "g fine p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2296
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2297
  presume "\<not> (\<exists>p. p tagged_division_of (cbox a b) \<and> g fine p) \<Longrightarrow> False"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2298
  then obtain p where "p tagged_division_of (cbox a b)" "g fine p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2299
    by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2300
  then show thesis ..
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2301
next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2302
  assume as: "\<not> (\<exists>p. p tagged_division_of (cbox a b) \<and> g fine p)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2303
  obtain x where x:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2304
      "x \<in> (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2305
      "\<And>e. 0 < e \<Longrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2306
        \<exists>c d.
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2307
          x \<in> cbox c d \<and>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2308
          cbox c d \<subseteq> ball x e \<and>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2309
          cbox c d \<subseteq> (cbox a b) \<and>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2310
          \<not> (\<exists>p. p tagged_division_of cbox c d \<and> g fine p)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2311
    apply (rule interval_bisection[of "\<lambda>s. \<exists>p. p tagged_division_of s \<and> g fine p", OF _ _ as])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2312
    apply (simp add: fine_def)
66192
e5b84854baa4 A few renamings and several tidied-up proofs
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
  2313
    apply (metis tagged_division_union fine_Un)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2314
    apply (auto simp: )
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2315
    done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2316
  obtain e where e: "e > 0" "ball x e \<subseteq> g x"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2317
    using gaugeD[OF assms, of x] unfolding open_contains_ball by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2318
  from x(2)[OF e(1)]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2319
  obtain c d where c_d: "x \<in> cbox c d"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2320
                        "cbox c d \<subseteq> ball x e"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2321
                        "cbox c d \<subseteq> cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2322
                        "\<not> (\<exists>p. p tagged_division_of cbox c d \<and> g fine p)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2323
    by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2324
  have "g fine {(x, cbox c d)}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2325
    unfolding fine_def using e using c_d(2) by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2326
  then show False
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2327
    using tagged_division_of_self[OF c_d(1)] using c_d by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2328
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2329
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2330
lemma fine_division_exists_real:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2331
  fixes a b :: real
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2332
  assumes "gauge g"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2333
  obtains p where "p tagged_division_of {a .. b}" "g fine p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2334
  by (metis assms box_real(2) fine_division_exists)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2335
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2336
subsection \<open>A technical lemma about "refinement" of division.\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2337
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2338
lemma tagged_division_finer:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2339
  fixes p :: "('a::euclidean_space \<times> ('a::euclidean_space set)) set"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2340
  assumes "p tagged_division_of (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2341
    and "gauge d"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2342
  obtains q where "q tagged_division_of (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2343
    and "d fine q"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2344
    and "\<forall>(x,k) \<in> p. k \<subseteq> d(x) \<longrightarrow> (x,k) \<in> q"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2345
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2346
  let ?P = "\<lambda>p. p tagged_partial_division_of (cbox a b) \<longrightarrow> gauge d \<longrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2347
    (\<exists>q. q tagged_division_of (\<Union>{k. \<exists>x. (x,k) \<in> p}) \<and> d fine q \<and>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2348
      (\<forall>(x,k) \<in> p. k \<subseteq> d(x) \<longrightarrow> (x,k) \<in> q))"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2349
  {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2350
    have *: "finite p" "p tagged_partial_division_of (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2351
      using assms(1)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2352
      unfolding tagged_division_of_def
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2353
      by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2354
    presume "\<And>p. finite p \<Longrightarrow> ?P p"
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  2355
    from this[rule_format,OF * assms(2)] 
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  2356
    obtain q where q: "q tagged_division_of \<Union>{k. \<exists>x. (x, k) \<in> p}" "d fine q" "(\<forall>(x, k)\<in>p. k \<subseteq> d x \<longrightarrow> (x, k) \<in> q)"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  2357
      by auto
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  2358
    with that[of q] show ?thesis
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  2359
      using assms(1) by auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2360
  }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2361
  fix p :: "('a::euclidean_space \<times> ('a::euclidean_space set)) set"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2362
  assume as: "finite p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2363
  show "?P p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2364
    apply rule
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2365
    apply rule
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2366
    using as
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2367
  proof (induct p)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2368
    case empty
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2369
    show ?case
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2370
      apply (rule_tac x="{}" in exI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2371
      unfolding fine_def
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2372
      apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2373
      done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2374
  next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2375
    case (insert xk p)
66314
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2376
    obtain x k where xk: "xk = (x, k)"
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2377
      using surj_pair[of xk] by metis 
66193
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2378
    obtain q1 where q1: "q1 tagged_division_of \<Union>{k. \<exists>x. (x, k) \<in> p}"
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2379
                and "d fine q1"
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2380
                and q1I: "\<And>x k. \<lbrakk>(x, k)\<in>p;  k \<subseteq> d x\<rbrakk> \<Longrightarrow> (x, k) \<in> q1"
66314
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2381
      using case_prodD tagged_partial_division_subset[OF insert(4) subset_insertI]
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2382
      by (metis (mono_tags, lifting) insert.hyps(3) insert.prems(2))
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2383
    have *: "\<Union>{l. \<exists>y. (y,l) \<in> insert xk p} = k \<union> \<Union>{l. \<exists>y. (y,l) \<in> p}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2384
      unfolding xk by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2385
    note p = tagged_partial_division_ofD[OF insert(4)]
66314
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2386
    obtain u v where uv: "k = cbox u v"
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2387
      using p(4)[unfolded xk, OF insertI1] by blast
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2388
    have "finite {k. \<exists>x. (x, k) \<in> p}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2389
      apply (rule finite_subset[of _ "snd ` p"])
66314
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2390
      using image_iff apply fastforce
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2391
      using insert.hyps(1) by blast
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2392
    then have int: "interior (cbox u v) \<inter> interior (\<Union>{k. \<exists>x. (x, k) \<in> p}) = {}"
66314
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2393
    proof (rule Int_interior_Union_intervals)
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2394
      show "open (interior (cbox u v))"
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2395
        by auto
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2396
      show "\<And>T. T \<in> {k. \<exists>x. (x, k) \<in> p} \<Longrightarrow> \<exists>a b. T = cbox a b"
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2397
        using p(4) by auto
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2398
      show "\<And>T. T \<in> {k. \<exists>x. (x, k) \<in> p} \<Longrightarrow> interior (cbox u v) \<inter> interior T = {}"
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2399
        by clarify (metis insert.hyps(2) insert_iff interior_cbox p(5) uv xk)
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2400
    qed
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2401
    show ?case
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2402
    proof (cases "cbox u v \<subseteq> d x")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2403
      case True
66314
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2404
      have "{(x, cbox u v)} tagged_division_of cbox u v"
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2405
        by (simp add: p(2) uv xk tagged_division_of_self)
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2406
      then have "{(x, cbox u v)} \<union> q1 tagged_division_of \<Union>{k. \<exists>x. (x, k) \<in> insert xk p}"
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2407
        unfolding * uv by (metis (no_types, lifting) int q1 tagged_division_union)
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2408
      with True show ?thesis
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2409
        apply (rule_tac x="{(x,cbox u v)} \<union> q1" in exI)
66314
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2410
        using \<open>d fine q1\<close> fine_def q1I uv xk apply fastforce
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2411
        done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2412
    next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2413
      case False
66314
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2414
      obtain q2 where q2: "q2 tagged_division_of cbox u v" "d fine q2"
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2415
        using fine_division_exists[OF assms(2)] by blast
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2416
      show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2417
        apply (rule_tac x="q2 \<union> q1" in exI)
66314
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2418
        apply (intro conjI)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2419
        unfolding * uv
66192
e5b84854baa4 A few renamings and several tidied-up proofs
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
  2420
        apply (rule tagged_division_union q2 q1 int fine_Un)+
66193
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2421
          apply (auto intro: q1 q2 fine_Un \<open>d fine q1\<close> simp add: False q1I uv xk)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2422
        done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2423
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2424
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2425
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2426
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2427
subsubsection \<open>Covering lemma\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2428
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2429
text\<open> Some technical lemmas used in the approximation results that follow. Proof of the covering
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2430
  lemma is an obvious multidimensional generalization of Lemma 3, p65 of Swartz's
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2431
  "Introduction to Gauge Integrals". \<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2432
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2433
proposition covering_lemma:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2434
  assumes "S \<subseteq> cbox a b" "box a b \<noteq> {}" "gauge g"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2435
  obtains \<D> where
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2436
    "countable \<D>"  "\<Union>\<D> \<subseteq> cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2437
    "\<And>K. K \<in> \<D> \<Longrightarrow> interior K \<noteq> {} \<and> (\<exists>c d. K = cbox c d)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2438
    "pairwise (\<lambda>A B. interior A \<inter> interior B = {}) \<D>"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2439
    "\<And>K. K \<in> \<D> \<Longrightarrow> \<exists>x \<in> S \<inter> K. K \<subseteq> g x"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2440
    "\<And>u v. cbox u v \<in> \<D> \<Longrightarrow> \<exists>n. \<forall>i \<in> Basis. v \<bullet> i - u \<bullet> i = (b \<bullet> i - a \<bullet> i) / 2^n"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2441
    "S \<subseteq> \<Union>\<D>"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2442
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2443
  have aibi: "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i" and normab: "0 < norm(b - a)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2444
    using \<open>box a b \<noteq> {}\<close> box_eq_empty box_sing by fastforce+
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2445
  let ?K0 = "\<lambda>(n, f::'a\<Rightarrow>nat).
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2446
                    cbox (\<Sum>i \<in> Basis. (a \<bullet> i + (f i / 2^n) * (b \<bullet> i - a \<bullet> i)) *\<^sub>R i)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2447
                         (\<Sum>i \<in> Basis. (a \<bullet> i + ((f i + 1) / 2^n) * (b \<bullet> i - a \<bullet> i)) *\<^sub>R i)"
64910
6108dddad9f0 more symbols via abbrevs;
wenzelm
parents: 64267
diff changeset
  2448
  let ?D0 = "?K0 ` (SIGMA n:UNIV. Pi\<^sub>E Basis (\<lambda>i::'a. lessThan (2^n)))"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2449
  obtain \<D>0 where count: "countable \<D>0"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2450
             and sub: "\<Union>\<D>0 \<subseteq> cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2451
             and int:  "\<And>K. K \<in> \<D>0 \<Longrightarrow> (interior K \<noteq> {}) \<and> (\<exists>c d. K = cbox c d)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2452
             and intdj: "\<And>A B. \<lbrakk>A \<in> \<D>0; B \<in> \<D>0\<rbrakk> \<Longrightarrow> A \<subseteq> B \<or> B \<subseteq> A \<or> interior A \<inter> interior B = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2453
             and SK: "\<And>x. x \<in> S \<Longrightarrow> \<exists>K \<in> \<D>0. x \<in> K \<and> K \<subseteq> g x"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2454
             and cbox: "\<And>u v. cbox u v \<in> \<D>0 \<Longrightarrow> \<exists>n. \<forall>i \<in> Basis. v \<bullet> i - u \<bullet> i = (b \<bullet> i - a \<bullet> i) / 2^n"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2455
             and fin: "\<And>K. K \<in> \<D>0 \<Longrightarrow> finite {L \<in> \<D>0. K \<subseteq> L}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2456
  proof
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2457
    show "countable ?D0"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2458
      by (simp add: countable_PiE)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2459
  next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2460
    show "\<Union>?D0 \<subseteq> cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2461
      apply (simp add: UN_subset_iff)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2462
      apply (intro conjI allI ballI subset_box_imp)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2463
       apply (simp add: divide_simps zero_le_mult_iff aibi)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2464
      apply (force simp: aibi scaling_mono nat_less_real_le dest: PiE_mem)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2465
      done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2466
  next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2467
    show "\<And>K. K \<in> ?D0 \<Longrightarrow> interior K \<noteq> {} \<and> (\<exists>c d. K = cbox c d)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2468
      using \<open>box a b \<noteq> {}\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2469
      by (clarsimp simp: box_eq_empty) (fastforce simp add: divide_simps dest: PiE_mem)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2470
  next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2471
    have realff: "(real w) * 2^m < (real v) * 2^n \<longleftrightarrow> w * 2^m < v * 2^n" for m n v w
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2472
      using of_nat_less_iff less_imp_of_nat_less by fastforce
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2473
    have *: "\<forall>v w. ?K0(m,v) \<subseteq> ?K0(n,w) \<or> ?K0(n,w) \<subseteq> ?K0(m,v) \<or> interior(?K0(m,v)) \<inter> interior(?K0(n,w)) = {}"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  2474
      for m n \<comment>\<open>The symmetry argument requires a single HOL formula\<close>
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2475
    proof (rule linorder_wlog [where a=m and b=n], intro allI impI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2476
      fix v w m and n::nat
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  2477
      assume "m \<le> n" \<comment>\<open>WLOG we can assume @{term"m \<le> n"}, when the first disjunct becomes impossible\<close>
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2478
      have "?K0(n,w) \<subseteq> ?K0(m,v) \<or> interior(?K0(m,v)) \<inter> interior(?K0(n,w)) = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2479
        apply (simp add: subset_box disjoint_interval)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2480
        apply (rule ccontr)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2481
        apply (clarsimp simp add: aibi mult_le_cancel_right divide_le_cancel not_less not_le)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2482
        apply (drule_tac x=i in bspec, assumption)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2483
        using \<open>m\<le>n\<close> realff [of _ _ "1+_"] realff [of "1+_"_ "1+_"]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2484
        apply (auto simp: divide_simps add.commute not_le nat_le_iff_add realff)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2485
        apply (simp add: power_add, metis (no_types, hide_lams) mult_Suc mult_less_cancel2 not_less_eq mult.assoc)+
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2486
        done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2487
      then show "?K0(m,v) \<subseteq> ?K0(n,w) \<or> ?K0(n,w) \<subseteq> ?K0(m,v) \<or> interior(?K0(m,v)) \<inter> interior(?K0(n,w)) = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2488
        by meson
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2489
    qed auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2490
    show "\<And>A B. \<lbrakk>A \<in> ?D0; B \<in> ?D0\<rbrakk> \<Longrightarrow> A \<subseteq> B \<or> B \<subseteq> A \<or> interior A \<inter> interior B = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2491
      apply (erule imageE SigmaE)+
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2492
      using * by simp
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2493
  next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2494
    show "\<exists>K \<in> ?D0. x \<in> K \<and> K \<subseteq> g x" if "x \<in> S" for x
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2495
    proof (simp only: bex_simps split_paired_Bex_Sigma)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2496
      show "\<exists>n. \<exists>f \<in> Basis \<rightarrow>\<^sub>E {..<2 ^ n}. x \<in> ?K0(n,f) \<and> ?K0(n,f) \<subseteq> g x"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2497
      proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2498
        obtain e where "0 < e"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2499
                   and e: "\<And>y. (\<And>i. i \<in> Basis \<Longrightarrow> \<bar>x \<bullet> i - y \<bullet> i\<bar> \<le> e) \<Longrightarrow> y \<in> g x"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2500
        proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2501
          have "x \<in> g x" "open (g x)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2502
            using \<open>gauge g\<close> by (auto simp: gauge_def)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2503
          then obtain \<epsilon> where "0 < \<epsilon>" and \<epsilon>: "ball x \<epsilon> \<subseteq> g x"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2504
            using openE by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2505
          have "norm (x - y) < \<epsilon>"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2506
               if "(\<And>i. i \<in> Basis \<Longrightarrow> \<bar>x \<bullet> i - y \<bullet> i\<bar> \<le> \<epsilon> / (2 * real DIM('a)))" for y
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2507
          proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2508
            have "norm (x - y) \<le> (\<Sum>i\<in>Basis. \<bar>x \<bullet> i - y \<bullet> i\<bar>)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
  2509
              by (metis (no_types, lifting) inner_diff_left norm_le_l1 sum.cong)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2510
            also have "... \<le> DIM('a) * (\<epsilon> / (2 * real DIM('a)))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
  2511
              by (meson sum_bounded_above that)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2512
            also have "... = \<epsilon> / 2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2513
              by (simp add: divide_simps)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2514
            also have "... < \<epsilon>"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2515
              by (simp add: \<open>0 < \<epsilon>\<close>)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2516
            finally show ?thesis .
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2517
          qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2518
          then show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2519
            by (rule_tac e = "\<epsilon> / 2 / DIM('a)" in that) (simp_all add:  \<open>0 < \<epsilon>\<close> dist_norm subsetD [OF \<epsilon>])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2520
        qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2521
        have xab: "x \<in> cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2522
          using \<open>x \<in> S\<close> \<open>S \<subseteq> cbox a b\<close> by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2523
        obtain n where n: "norm (b - a) / 2^n < e"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2524
          using real_arch_pow_inv [of "e / norm(b - a)" "1/2"] normab \<open>0 < e\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2525
          by (auto simp: divide_simps)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2526
        then have "norm (b - a) < e * 2^n"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2527
          by (auto simp: divide_simps)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2528
        then have bai: "b \<bullet> i - a \<bullet> i < e * 2 ^ n" if "i \<in> Basis" for i
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2529
        proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2530
          have "b \<bullet> i - a \<bullet> i \<le> norm (b - a)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2531
            by (metis abs_of_nonneg dual_order.trans inner_diff_left linear norm_ge_zero Basis_le_norm that)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2532
          also have "... < e * 2 ^ n"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2533
            using \<open>norm (b - a) < e * 2 ^ n\<close> by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2534
          finally show ?thesis .
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2535
        qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2536
        have D: "(a + n \<le> x \<and> x \<le> a + m) \<Longrightarrow> (a + n \<le> y \<and> y \<le> a + m) \<Longrightarrow> abs(x - y) \<le> m - n"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2537
                 for a m n x and y::real
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2538
          by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2539
        have "\<forall>i\<in>Basis. \<exists>k<2 ^ n. (a \<bullet> i + real k * (b \<bullet> i - a \<bullet> i) / 2 ^ n \<le> x \<bullet> i \<and>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2540
               x \<bullet> i \<le> a \<bullet> i + (real k + 1) * (b \<bullet> i - a \<bullet> i) / 2 ^ n)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2541
        proof
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2542
          fix i::'a assume "i \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2543
          consider "x \<bullet> i = b \<bullet> i" | "x \<bullet> i < b \<bullet> i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2544
            using \<open>i \<in> Basis\<close> mem_box(2) xab by force
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2545
          then show "\<exists>k<2 ^ n. (a \<bullet> i + real k * (b \<bullet> i - a \<bullet> i) / 2 ^ n \<le> x \<bullet> i \<and>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2546
                          x \<bullet> i \<le> a \<bullet> i + (real k + 1) * (b \<bullet> i - a \<bullet> i) / 2 ^ n)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2547
          proof cases
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2548
            case 1 then show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2549
              by (rule_tac x = "2^n - 1" in exI) (auto simp: algebra_simps divide_simps of_nat_diff \<open>i \<in> Basis\<close> aibi)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2550
          next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2551
            case 2
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2552
            then have abi_less: "a \<bullet> i < b \<bullet> i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2553
              using \<open>i \<in> Basis\<close> xab by (auto simp: mem_box)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2554
            let ?k = "nat \<lfloor>2 ^ n * (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i)\<rfloor>"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2555
            show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2556
            proof (intro exI conjI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2557
              show "?k < 2 ^ n"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2558
                using aibi xab \<open>i \<in> Basis\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2559
                by (force simp: nat_less_iff floor_less_iff divide_simps 2 mem_box)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2560
            next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2561
              have "a \<bullet> i + real ?k * (b \<bullet> i - a \<bullet> i) / 2 ^ n \<le>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2562
                    a \<bullet> i + (2 ^ n * (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i)) * (b \<bullet> i - a \<bullet> i) / 2 ^ n"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2563
                apply (intro add_left_mono mult_right_mono divide_right_mono of_nat_floor)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2564
                using aibi [OF \<open>i \<in> Basis\<close>] xab 2
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2565
                  apply (simp_all add: \<open>i \<in> Basis\<close> mem_box divide_simps)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2566
                done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2567
              also have "... = x \<bullet> i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2568
                using abi_less by (simp add: divide_simps)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2569
              finally show "a \<bullet> i + real ?k * (b \<bullet> i - a \<bullet> i) / 2 ^ n \<le> x \<bullet> i" .
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2570
            next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2571
              have "x \<bullet> i \<le> a \<bullet> i + (2 ^ n * (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i)) * (b \<bullet> i - a \<bullet> i) / 2 ^ n"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2572
                using abi_less by (simp add: divide_simps algebra_simps)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2573
              also have "... \<le> a \<bullet> i + (real ?k + 1) * (b \<bullet> i - a \<bullet> i) / 2 ^ n"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2574
                apply (intro add_left_mono mult_right_mono divide_right_mono of_nat_floor)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2575
                using aibi [OF \<open>i \<in> Basis\<close>] xab
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2576
                  apply (auto simp: \<open>i \<in> Basis\<close> mem_box divide_simps)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2577
                done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2578
              finally show "x \<bullet> i \<le> a \<bullet> i + (real ?k + 1) * (b \<bullet> i - a \<bullet> i) / 2 ^ n" .
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2579
            qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2580
          qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2581
        qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2582
        then have "\<exists>f\<in>Basis \<rightarrow>\<^sub>E {..<2 ^ n}. x \<in> ?K0(n,f)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2583
          apply (simp add: mem_box Bex_def)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2584
          apply (clarify dest!: bchoice)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2585
          apply (rule_tac x="restrict f Basis" in exI, simp)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2586
          done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2587
        moreover have "\<And>f. x \<in> ?K0(n,f) \<Longrightarrow> ?K0(n,f) \<subseteq> g x"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2588
          apply (clarsimp simp add: mem_box)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2589
          apply (rule e)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2590
          apply (drule bspec D, assumption)+
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2591
          apply (erule order_trans)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2592
          apply (simp add: divide_simps)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2593
          using bai by (force simp: algebra_simps)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2594
        ultimately show ?thesis by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2595
      qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2596
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2597
  next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2598
    show "\<And>u v. cbox u v \<in> ?D0 \<Longrightarrow> \<exists>n. \<forall>i \<in> Basis. v \<bullet> i - u \<bullet> i = (b \<bullet> i - a \<bullet> i) / 2^n"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2599
      by (force simp: eq_cbox box_eq_empty field_simps dest!: aibi)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2600
  next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2601
    obtain j::'a where "j \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2602
      using nonempty_Basis by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2603
    have "finite {L \<in> ?D0. ?K0(n,f) \<subseteq> L}" if "f \<in> Basis \<rightarrow>\<^sub>E {..<2 ^ n}" for n f
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2604
    proof (rule finite_subset)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2605
      let ?B = "(\<lambda>(n, f::'a\<Rightarrow>nat). cbox (\<Sum>i\<in>Basis. (a \<bullet> i + (f i) / 2^n * (b \<bullet> i - a \<bullet> i)) *\<^sub>R i)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2606
                                        (\<Sum>i\<in>Basis. (a \<bullet> i + ((f i) + 1) / 2^n * (b \<bullet> i - a \<bullet> i)) *\<^sub>R i))
64910
6108dddad9f0 more symbols via abbrevs;
wenzelm
parents: 64267
diff changeset
  2607
                ` (SIGMA m:atMost n. Pi\<^sub>E Basis (\<lambda>i::'a. lessThan (2^m)))"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2608
      have "?K0(m,g) \<in> ?B" if "g \<in> Basis \<rightarrow>\<^sub>E {..<2 ^ m}" "?K0(n,f) \<subseteq> ?K0(m,g)" for m g
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2609
      proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2610
        have dd: "w / m \<le> v / n \<and> (v+1) / n \<le> (w+1) / m
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2611
                  \<Longrightarrow> inverse n \<le> inverse m" for w m v n::real
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2612
          by (auto simp: divide_simps algebra_simps)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2613
        have bjaj: "b \<bullet> j - a \<bullet> j > 0"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2614
          using \<open>j \<in> Basis\<close> \<open>box a b \<noteq> {}\<close> box_eq_empty(1) by fastforce
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2615
        have "((g j) / 2 ^ m) * (b \<bullet> j - a \<bullet> j) \<le> ((f j) / 2 ^ n) * (b \<bullet> j - a \<bullet> j) \<and>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2616
              (((f j) + 1) / 2 ^ n) * (b \<bullet> j - a \<bullet> j) \<le> (((g j) + 1) / 2 ^ m) * (b \<bullet> j - a \<bullet> j)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2617
          using that \<open>j \<in> Basis\<close> by (simp add: subset_box algebra_simps divide_simps aibi)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2618
        then have "((g j) / 2 ^ m) \<le> ((f j) / 2 ^ n) \<and>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2619
          ((real(f j) + 1) / 2 ^ n) \<le> ((real(g j) + 1) / 2 ^ m)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2620
          by (metis bjaj mult.commute of_nat_1 of_nat_add real_mult_le_cancel_iff2)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2621
        then have "inverse (2^n) \<le> (inverse (2^m) :: real)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2622
          by (rule dd)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2623
        then have "m \<le> n"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2624
          by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2625
        show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2626
          by (rule imageI) (simp add: \<open>m \<le> n\<close> that)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2627
      qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2628
      then show "{L \<in> ?D0. ?K0(n,f) \<subseteq> L} \<subseteq> ?B"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2629
        by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2630
      show "finite ?B"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2631
        by (intro finite_imageI finite_SigmaI finite_atMost finite_lessThan finite_PiE finite_Basis)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2632
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2633
    then show "finite {L \<in> ?D0. K \<subseteq> L}" if "K \<in> ?D0" for K
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2634
      using that by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2635
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2636
  let ?D1 = "{K \<in> \<D>0. \<exists>x \<in> S \<inter> K. K \<subseteq> g x}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2637
  obtain \<D> where count: "countable \<D>"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2638
             and sub: "\<Union>\<D> \<subseteq> cbox a b"  "S \<subseteq> \<Union>\<D>"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2639
             and int:  "\<And>K. K \<in> \<D> \<Longrightarrow> (interior K \<noteq> {}) \<and> (\<exists>c d. K = cbox c d)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2640
             and intdj: "\<And>A B. \<lbrakk>A \<in> \<D>; B \<in> \<D>\<rbrakk> \<Longrightarrow> A \<subseteq> B \<or> B \<subseteq> A \<or> interior A \<inter> interior B = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2641
             and SK: "\<And>K. K \<in> \<D> \<Longrightarrow> \<exists>x. x \<in> S \<inter> K \<and> K \<subseteq> g x"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2642
             and cbox: "\<And>u v. cbox u v \<in> \<D> \<Longrightarrow> \<exists>n. \<forall>i \<in> Basis. v \<bullet> i - u \<bullet> i = (b \<bullet> i - a \<bullet> i) / 2^n"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2643
             and fin: "\<And>K. K \<in> \<D> \<Longrightarrow> finite {L. L \<in> \<D> \<and> K \<subseteq> L}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2644
  proof
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2645
    show "countable ?D1" using count countable_subset
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2646
      by (simp add: count countable_subset)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2647
    show "\<Union>?D1 \<subseteq> cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2648
      using sub by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2649
    show "S \<subseteq> \<Union>?D1"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2650
      using SK by (force simp:)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2651
    show "\<And>K. K \<in> ?D1 \<Longrightarrow> (interior K \<noteq> {}) \<and> (\<exists>c d. K = cbox c d)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2652
      using int by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2653
    show "\<And>A B. \<lbrakk>A \<in> ?D1; B \<in> ?D1\<rbrakk> \<Longrightarrow> A \<subseteq> B \<or> B \<subseteq> A \<or> interior A \<inter> interior B = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2654
      using intdj by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2655
    show "\<And>K. K \<in> ?D1 \<Longrightarrow> \<exists>x. x \<in> S \<inter> K \<and> K \<subseteq> g x"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2656
      by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2657
    show "\<And>u v. cbox u v \<in> ?D1 \<Longrightarrow> \<exists>n. \<forall>i \<in> Basis. v \<bullet> i - u \<bullet> i = (b \<bullet> i - a \<bullet> i) / 2^n"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2658
      using cbox by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2659
    show "\<And>K. K \<in> ?D1 \<Longrightarrow> finite {L. L \<in> ?D1 \<and> K \<subseteq> L}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2660
      using fin by simp (metis (mono_tags, lifting) Collect_mono rev_finite_subset)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2661
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2662
  let ?\<D> = "{K \<in> \<D>. \<forall>K'. K' \<in> \<D> \<and> K \<noteq> K' \<longrightarrow> ~(K \<subseteq> K')}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2663
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2664
  proof (rule that)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2665
    show "countable ?\<D>"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2666
      by (blast intro: countable_subset [OF _ count])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2667
    show "\<Union>?\<D> \<subseteq> cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2668
      using sub by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2669
    show "S \<subseteq> \<Union>?\<D>"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2670
    proof clarsimp
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2671
      fix x
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2672
      assume "x \<in> S"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2673
      then obtain X where "x \<in> X" "X \<in> \<D>" using \<open>S \<subseteq> \<Union>\<D>\<close> by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2674
      let ?R = "{(K,L). K \<in> \<D> \<and> L \<in> \<D> \<and> L \<subset> K}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2675
      have irrR: "irrefl ?R" by (force simp: irrefl_def)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2676
      have traR: "trans ?R" by (force simp: trans_def)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2677
      have finR: "\<And>x. finite {y. (y, x) \<in> ?R}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2678
        by simp (metis (mono_tags, lifting) fin \<open>X \<in> \<D>\<close> finite_subset mem_Collect_eq psubset_imp_subset subsetI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2679
      have "{X \<in> \<D>. x \<in> X} \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2680
        using \<open>X \<in> \<D>\<close> \<open>x \<in> X\<close> by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2681
      then obtain Y where "Y \<in> {X \<in> \<D>. x \<in> X}" "\<And>Y'. (Y', Y) \<in> ?R \<Longrightarrow> Y' \<notin> {X \<in> \<D>. x \<in> X}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2682
        by (rule wfE_min' [OF wf_finite_segments [OF irrR traR finR]]) blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2683
      then show "\<exists>Y. Y \<in> \<D> \<and> (\<forall>K'. K' \<in> \<D> \<and> Y \<noteq> K' \<longrightarrow> \<not> Y \<subseteq> K') \<and> x \<in> Y"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2684
        by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2685
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2686
    show "\<And>K. K \<in> ?\<D> \<Longrightarrow> interior K \<noteq> {} \<and> (\<exists>c d. K = cbox c d)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2687
      by (blast intro: dest: int)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2688
    show "pairwise (\<lambda>A B. interior A \<inter> interior B = {}) ?\<D>"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2689
      using intdj by (simp add: pairwise_def) metis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2690
    show "\<And>K. K \<in> ?\<D> \<Longrightarrow> \<exists>x \<in> S \<inter> K. K \<subseteq> g x"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2691
      using SK by force
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2692
    show "\<And>u v. cbox u v \<in> ?\<D> \<Longrightarrow> \<exists>n. \<forall>i\<in>Basis. v \<bullet> i - u \<bullet> i = (b \<bullet> i - a \<bullet> i) / 2^n"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2693
      using cbox by force
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2694
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2695
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2696
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2697
subsection \<open>Division filter\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2698
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2699
text \<open>Divisions over all gauges towards finer divisions.\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2700
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2701
definition division_filter :: "'a::euclidean_space set \<Rightarrow> ('a \<times> 'a set) set filter"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2702
  where "division_filter s = (INF g:{g. gauge g}. principal {p. p tagged_division_of s \<and> g fine p})"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2703
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2704
lemma eventually_division_filter:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2705
  "(\<forall>\<^sub>F p in division_filter s. P p) \<longleftrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2706
    (\<exists>g. gauge g \<and> (\<forall>p. p tagged_division_of s \<and> g fine p \<longrightarrow> P p))"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2707
  unfolding division_filter_def
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2708
proof (subst eventually_INF_base; clarsimp)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2709
  fix g1 g2 :: "'a \<Rightarrow> 'a set" show "gauge g1 \<Longrightarrow> gauge g2 \<Longrightarrow> \<exists>x. gauge x \<and>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2710
    {p. p tagged_division_of s \<and> x fine p} \<subseteq> {p. p tagged_division_of s \<and> g1 fine p} \<and>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2711
    {p. p tagged_division_of s \<and> x fine p} \<subseteq> {p. p tagged_division_of s \<and> g2 fine p}"
66192
e5b84854baa4 A few renamings and several tidied-up proofs
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
  2712
    by (intro exI[of _ "\<lambda>x. g1 x \<inter> g2 x"]) (auto simp: fine_Int)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2713
qed (auto simp: eventually_principal)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2714
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2715
lemma division_filter_not_empty: "division_filter (cbox a b) \<noteq> bot"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2716
  unfolding trivial_limit_def eventually_division_filter
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2717
  by (auto elim: fine_division_exists)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2718
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2719
lemma eventually_division_filter_tagged_division:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2720
  "eventually (\<lambda>p. p tagged_division_of s) (division_filter s)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2721
  unfolding eventually_division_filter by (intro exI[of _ "\<lambda>x. ball x 1"]) auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2722
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
  2723
end