src/HOL/Analysis/Tagged_Division.thy
author paulson <lp15@cam.ac.uk>
Thu, 24 Aug 2017 12:45:46 +0100
changeset 66498 97fc319d6089
parent 66492 d7206afe2d28
parent 66497 18a6478a574c
child 66703 61bf958fa1c1
permissions -rw-r--r--
Merge (non-trivial)
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
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
    13
term comm_monoid
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
    14
66294
0442b3f45556 refactored some HORRIBLE integration proofs
paulson <lp15@cam.ac.uk>
parents: 66193
diff changeset
    15
lemma sum_Sigma_product:
0442b3f45556 refactored some HORRIBLE integration proofs
paulson <lp15@cam.ac.uk>
parents: 66193
diff changeset
    16
  assumes "finite S"
0442b3f45556 refactored some HORRIBLE integration proofs
paulson <lp15@cam.ac.uk>
parents: 66193
diff changeset
    17
    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
    18
  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
    19
  using assms
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    20
proof induct
66294
0442b3f45556 refactored some HORRIBLE integration proofs
paulson <lp15@cam.ac.uk>
parents: 66193
diff changeset
    21
  case empty
0442b3f45556 refactored some HORRIBLE integration proofs
paulson <lp15@cam.ac.uk>
parents: 66193
diff changeset
    22
  then show ?case
0442b3f45556 refactored some HORRIBLE integration proofs
paulson <lp15@cam.ac.uk>
parents: 66193
diff changeset
    23
    by simp
0442b3f45556 refactored some HORRIBLE integration proofs
paulson <lp15@cam.ac.uk>
parents: 66193
diff changeset
    24
next
0442b3f45556 refactored some HORRIBLE integration proofs
paulson <lp15@cam.ac.uk>
parents: 66193
diff changeset
    25
  case (insert a S)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    26
  show ?case
66294
0442b3f45556 refactored some HORRIBLE integration proofs
paulson <lp15@cam.ac.uk>
parents: 66193
diff changeset
    27
    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
    28
qed
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    29
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    30
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
    31
  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
    32
  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
    33
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    34
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    35
subsection \<open>Sundries\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    36
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    37
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    38
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
    39
lemma wf_finite_segments:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    40
  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
    41
  shows "wf (r)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    42
  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
    43
  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
    44
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    45
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
    46
lemma scaling_mono:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    47
  fixes u::"'a::linordered_field"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    48
  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
    49
    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
    50
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    51
  have "r/s \<le> 1" using assms
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    52
    using divide_le_eq_1 by fastforce
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    53
  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
    54
    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
    55
  then show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    56
    by (simp add: field_simps)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    57
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    58
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    59
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
    60
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    61
lemma interior_subset_union_intervals:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    62
  assumes "i = cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    63
    and "j = cbox c d"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    64
    and "interior j \<noteq> {}"
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
    65
    and "i \<subseteq> j \<union> S"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    66
    and "interior i \<inter> interior j = {}"
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
    67
  shows "interior i \<subseteq> interior S"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    68
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    69
  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
    70
     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
    71
     unfolding interior_cbox by auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    72
  moreover
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
    73
  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
    74
    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
    75
    using assms by auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    76
  ultimately
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    77
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    78
    unfolding assms interior_cbox
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    79
      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
    80
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    81
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    82
lemma interior_Union_subset_cbox:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    83
  assumes "finite f"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    84
  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
    85
    and t: "closed t"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    86
  shows "interior (\<Union>f) \<subseteq> t"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    87
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    88
  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
    89
    using f by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    90
  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
    91
  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
    92
    using \<open>finite f\<close> by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    93
  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
    94
  proof (induction E rule: finite_subset_induct')
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    95
    case (insert s f')
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
    96
    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
    97
      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
    98
    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
    99
      using insert.hyps by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   100
    finally show ?case
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   101
      by (simp add: insert.IH)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   102
  qed simp
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   103
  also have "\<dots> \<subseteq> \<Union>(f - E)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   104
    by (rule interior_subset)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   105
  also have "\<dots> \<subseteq> t"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   106
  proof (rule Union_least)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   107
    fix s assume "s \<in> f - E"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   108
    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
   109
      by (fastforce simp: E_def)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   110
    have "closure (interior s) \<subseteq> closure t"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   111
      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
   112
    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
   113
      by simp
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   114
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   115
  finally show ?thesis .
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   116
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   117
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   118
lemma Int_interior_Union_intervals:
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   119
    "\<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
   120
    \<Longrightarrow> S \<inter> interior (\<Union>\<F>) = {}"
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   121
  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
   122
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   123
lemma interval_split:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   124
  fixes a :: "'a::euclidean_space"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   125
  assumes "k \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   126
  shows
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   127
    "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
   128
    "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
   129
  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
   130
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   131
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
   132
  by (simp add: box_ne_empty)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   133
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   134
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
   135
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   136
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
   137
  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
   138
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   139
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
   140
  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
   141
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   142
lemma interval_upperbound[simp]:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   143
  "\<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
   144
    interval_upperbound (cbox a b) = (b::'a::euclidean_space)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
   145
  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
   146
  by (safe intro!: cSup_eq) auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   147
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   148
lemma interval_lowerbound[simp]:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   149
  "\<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
   150
    interval_lowerbound (cbox a b) = (a::'a::euclidean_space)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
   151
  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
   152
  by (safe intro!: cInf_eq) auto
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
lemmas interval_bounds = interval_upperbound interval_lowerbound
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   155
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   156
lemma
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   157
  fixes X::"real set"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   158
  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
   159
    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
   160
  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
   161
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   162
lemma interval_bounds'[simp]:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   163
  assumes "cbox a b \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   164
  shows "interval_upperbound (cbox a b) = b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   165
    and "interval_lowerbound (cbox a b) = a"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   166
  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
   167
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   168
lemma interval_upperbound_Times:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   169
  assumes "A \<noteq> {}" and "B \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   170
  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
   171
proof-
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   172
  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
   173
  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
   174
      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
   175
  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
   176
  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
   177
      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
   178
  ultimately show ?thesis unfolding interval_upperbound_def
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
   179
      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
   180
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   181
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   182
lemma interval_lowerbound_Times:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   183
  assumes "A \<noteq> {}" and "B \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   184
  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
   185
proof-
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   186
  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
   187
  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
   188
      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
   189
  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
   190
  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
   191
      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
   192
  ultimately show ?thesis unfolding interval_lowerbound_def
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
   193
      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
   194
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   195
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   196
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
   197
66317
a9bb833ee971 more tidying up
paulson
parents: 66314
diff changeset
   198
definition "gauge \<gamma> \<longleftrightarrow> (\<forall>x. x \<in> \<gamma> x \<and> open (\<gamma> x))"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   199
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   200
lemma gaugeI:
66317
a9bb833ee971 more tidying up
paulson
parents: 66314
diff changeset
   201
  assumes "\<And>x. x \<in> \<gamma> x"
a9bb833ee971 more tidying up
paulson
parents: 66314
diff changeset
   202
    and "\<And>x. open (\<gamma> x)"
a9bb833ee971 more tidying up
paulson
parents: 66314
diff changeset
   203
  shows "gauge \<gamma>"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   204
  using assms unfolding gauge_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   205
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   206
lemma gaugeD[dest]:
66317
a9bb833ee971 more tidying up
paulson
parents: 66314
diff changeset
   207
  assumes "gauge \<gamma>"
a9bb833ee971 more tidying up
paulson
parents: 66314
diff changeset
   208
  shows "x \<in> \<gamma> x"
a9bb833ee971 more tidying up
paulson
parents: 66314
diff changeset
   209
    and "open (\<gamma> x)"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   210
  using assms unfolding gauge_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   211
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   212
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
   213
  unfolding gauge_def by auto
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 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
   216
  unfolding gauge_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   217
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   218
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
   219
  by (rule gauge_ball) auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   220
66317
a9bb833ee971 more tidying up
paulson
parents: 66314
diff changeset
   221
lemma gauge_Int[intro]: "gauge \<gamma>1 \<Longrightarrow> gauge \<gamma>2 \<Longrightarrow> gauge (\<lambda>x. \<gamma>1 x \<inter> \<gamma>2 x)"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   222
  unfolding gauge_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   223
66164
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 66154
diff changeset
   224
lemma gauge_reflect:
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   225
  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
   226
  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
   227
  using equation_minus_iff
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 66154
diff changeset
   228
  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
   229
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 66154
diff changeset
   230
lemma gauge_Inter:
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   231
  assumes "finite S"
66317
a9bb833ee971 more tidying up
paulson
parents: 66314
diff changeset
   232
    and "\<And>u. u\<in>S \<Longrightarrow> gauge (f u)"
a9bb833ee971 more tidying up
paulson
parents: 66314
diff changeset
   233
  shows "gauge (\<lambda>x. \<Inter>{f u x | u. u \<in> S})"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   234
proof -
66317
a9bb833ee971 more tidying up
paulson
parents: 66314
diff changeset
   235
  have *: "\<And>x. {f u x |u. u \<in> S} = (\<lambda>u. f u x) ` S"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   236
    by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   237
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   238
    unfolding gauge_def unfolding *
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   239
    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
   240
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   241
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   242
lemma gauge_existence_lemma:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   243
  "(\<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
   244
  by (metis zero_less_one)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   245
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   246
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
   247
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   248
lemma gauge_modify:
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   249
  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
   250
  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
   251
  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
   252
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   253
subsection \<open>Divisions.\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   254
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   255
definition division_of (infixl "division'_of" 40)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   256
where
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   257
  "s division_of i \<longleftrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   258
    finite s \<and>
66113
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
   259
    (\<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
   260
    (\<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
   261
    (\<Union>s = i)"
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 division_ofD[dest]:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   264
  assumes "s division_of i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   265
  shows "finite s"
66113
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
   266
    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
   267
    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
   268
    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
   269
    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
   270
    and "\<Union>s = i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   271
  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
   272
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   273
lemma division_ofI:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   274
  assumes "finite s"
66113
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
   275
    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
   276
    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
   277
    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
   278
    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
   279
    and "\<Union>s = i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   280
  shows "s division_of i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   281
  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
   282
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   283
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
   284
  by auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   285
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   286
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
   287
  unfolding division_of_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   288
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   289
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
   290
  unfolding division_of_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   291
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   292
lemma division_of_sing[simp]:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   293
  "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
   294
  (is "?l = ?r")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   295
proof
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   296
  assume ?r
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   297
  moreover
66113
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
   298
  { fix K
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
   299
    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
   300
    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
   301
      apply (rule_tac x=a in exI)+
66317
a9bb833ee971 more tidying up
paulson
parents: 66314
diff changeset
   302
      apply force
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   303
      done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   304
  }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   305
  ultimately show ?l
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   306
    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
   307
next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   308
  assume ?l
66317
a9bb833ee971 more tidying up
paulson
parents: 66314
diff changeset
   309
  have "x = {a}" if  "x \<in> s" for x
a9bb833ee971 more tidying up
paulson
parents: 66314
diff changeset
   310
    by (metis \<open>s division_of cbox a a\<close> cbox_sing division_ofD(2) division_ofD(3) subset_singletonD that)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   311
  moreover have "s \<noteq> {}"
66193
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
   312
    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
   313
  ultimately show ?r
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   314
    unfolding cbox_sing by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   315
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   316
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   317
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
   318
  unfolding division_of_trivial by auto
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
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
   321
  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
   322
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   323
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
   324
  unfolding division_of_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   325
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   326
lemma forall_in_division:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   327
  "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
   328
  unfolding division_of_def by fastforce
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   329
66296
33a47f2d9edc New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
paulson <lp15@cam.ac.uk>
parents: 66295
diff changeset
   330
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
   331
  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
   332
  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
   333
  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
   334
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   335
lemma division_of_subset:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   336
  assumes "p division_of (\<Union>p)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   337
    and "q \<subseteq> p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   338
  shows "q division_of (\<Union>q)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   339
proof (rule division_ofI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   340
  note * = division_ofD[OF assms(1)]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   341
  show "finite q"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   342
    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
   343
  {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   344
    fix k
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   345
    assume "k \<in> q"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   346
    then have kp: "k \<in> p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   347
      using assms(2) by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   348
    show "k \<subseteq> \<Union>q"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   349
      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
   350
    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
   351
      using *(4)[OF kp] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   352
    show "k \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   353
      using *(3)[OF kp] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   354
  }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   355
  fix k1 k2
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   356
  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
   357
  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
   358
    using assms(2) by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   359
  show "interior k1 \<inter> interior k2 = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   360
    using *(5)[OF **] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   361
qed auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   362
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   363
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
   364
  unfolding division_of_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   365
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   366
lemma division_inter:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   367
  fixes s1 s2 :: "'a::euclidean_space set"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   368
  assumes "p1 division_of s1"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   369
    and "p2 division_of s2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   370
  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
   371
  (is "?A' division_of _")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   372
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   373
  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
   374
  have *: "?A' = ?A" by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   375
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   376
    unfolding *
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   377
  proof (rule division_ofI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   378
    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
   379
      by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   380
    moreover have "finite (p1 \<times> p2)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   381
      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
   382
    ultimately show "finite ?A" by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   383
    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
   384
      by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   385
    show "\<Union>?A = s1 \<inter> s2"
66317
a9bb833ee971 more tidying up
paulson
parents: 66314
diff changeset
   386
      unfolding * 
a9bb833ee971 more tidying up
paulson
parents: 66314
diff changeset
   387
      using division_ofD(6)[OF assms(1)] and division_ofD(6)[OF assms(2)] by auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   388
    {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   389
      fix k
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   390
      assume "k \<in> ?A"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   391
      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
   392
        by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   393
      then show "k \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   394
        by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   395
      show "k \<subseteq> s1 \<inter> s2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   396
        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
   397
        unfolding k by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   398
      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
   399
        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
   400
      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
   401
        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
   402
      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
   403
        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
   404
    }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   405
    fix k1 k2
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   406
    assume "k1 \<in> ?A"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   407
    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
   408
      by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   409
    assume "k2 \<in> ?A"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   410
    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
   411
      by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   412
    assume "k1 \<noteq> k2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   413
    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
   414
      unfolding k1 k2 by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   415
    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
   416
      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
   417
      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
   418
      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
   419
    show "interior k1 \<inter> interior k2 = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   420
      unfolding k1 k2
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   421
      apply (rule *)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   422
      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
   423
      done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   424
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   425
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   426
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   427
lemma division_inter_1:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   428
  assumes "d division_of i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   429
    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
   430
  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
   431
proof (cases "cbox a b = {}")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   432
  case True
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   433
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   434
    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
   435
next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   436
  case False
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   437
  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
   438
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   439
    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
   440
    unfolding * by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   441
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   442
66317
a9bb833ee971 more tidying up
paulson
parents: 66314
diff changeset
   443
lemma elementary_Int:
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   444
  fixes s t :: "'a::euclidean_space set"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   445
  assumes "p1 division_of s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   446
    and "p2 division_of t"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   447
  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
   448
using assms division_inter by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   449
66317
a9bb833ee971 more tidying up
paulson
parents: 66314
diff changeset
   450
lemma elementary_Inter:
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   451
  assumes "finite f"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   452
    and "f \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   453
    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
   454
  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
   455
  using assms
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   456
proof (induct f rule: finite_induct)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   457
  case (insert x f)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   458
  show ?case
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   459
  proof (cases "f = {}")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   460
    case True
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   461
    then show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   462
      unfolding True using insert by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   463
  next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   464
    case False
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   465
    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
   466
      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
   467
    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
   468
      using insert(5)[rule_format,OF insertI1] ..
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   469
    ultimately show ?thesis
66317
a9bb833ee971 more tidying up
paulson
parents: 66314
diff changeset
   470
      by (simp add: elementary_Int Inter_insert)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   471
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   472
qed auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   473
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   474
lemma division_disjoint_union:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   475
  assumes "p1 division_of s1"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   476
    and "p2 division_of s2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   477
    and "interior s1 \<inter> interior s2 = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   478
  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
   479
proof (rule division_ofI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   480
  note d1 = division_ofD[OF assms(1)]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   481
  note d2 = division_ofD[OF assms(2)]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   482
  show "finite (p1 \<union> p2)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   483
    using d1(1) d2(1) by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   484
  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
   485
    using d1(6) d2(6) by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   486
  {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   487
    fix k1 k2
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   488
    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
   489
    moreover
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   490
    let ?g="interior k1 \<inter> interior k2 = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   491
    {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   492
      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
   493
      have ?g
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   494
        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
   495
        using assms(3) by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   496
    }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   497
    moreover
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   498
    {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   499
      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
   500
      have ?g
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   501
        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
   502
        using assms(3) by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   503
    }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   504
    ultimately show ?g
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   505
      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
   506
  }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   507
  fix k
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   508
  assume k: "k \<in> p1 \<union> p2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   509
  show "k \<subseteq> s1 \<union> s2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   510
    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
   511
  show "k \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   512
    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
   513
  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
   514
    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
   515
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   516
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   517
lemma partial_division_extend_1:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   518
  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
   519
  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
   520
    and nonempty: "cbox c d \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   521
  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
   522
proof
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   523
  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
   524
    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
   525
  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
   526
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   527
  show "cbox c d \<in> p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   528
    unfolding p_def
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   529
    by (auto simp add: box_eq_empty cbox_def intro!: image_eqI[where x="\<lambda>(i::'a)\<in>Basis. (c, d)"])
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
   530
  have ord: "a \<bullet> i \<le> c \<bullet> i" "c \<bullet> i \<le> d \<bullet> i" "d \<bullet> i \<le> b \<bullet> i" if "i \<in> Basis" for i
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
   531
    using incl nonempty that
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
   532
    unfolding box_eq_empty subset_box by (auto simp: not_le)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   533
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   534
  show "p division_of (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   535
  proof (rule division_ofI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   536
    show "finite p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   537
      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
   538
    {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   539
      fix k
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   540
      assume "k \<in> p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   541
      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
   542
        by (auto simp: p_def)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   543
      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
   544
        by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   545
      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
   546
      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
   547
        fix i :: 'a
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   548
        assume i: "i \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   549
        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
   550
          by (auto simp: PiE_iff)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   551
        with i ord[of i]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   552
        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
   553
          by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   554
      qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   555
      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
   556
        by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   557
      {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   558
        fix l
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   559
        assume "l \<in> p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   560
        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
   561
          by (auto simp: p_def)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   562
        assume "l \<noteq> k"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   563
        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
   564
        proof (rule ccontr)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   565
          assume "\<not> ?thesis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   566
          with f g have "f = g"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   567
            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
   568
          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
   569
            by (simp add: l k)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   570
        qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   571
        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
   572
        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
   573
                  "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
   574
          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
   575
        with * ord[of i] show "interior l \<inter> interior k = {}"
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
   576
          by (auto simp add: l k disjoint_interval intro!: bexI[of _ i])
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   577
      }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   578
      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
   579
    }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   580
    moreover
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   581
    {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   582
      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
   583
      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
   584
      proof
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   585
        fix i :: 'a
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   586
        assume "i \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   587
        with x ord[of i]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   588
        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
   589
            (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
   590
          by (auto simp: cbox_def)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   591
        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
   592
          by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   593
      qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   594
      then obtain f where
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   595
        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
   596
        unfolding bchoice_iff ..
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   597
      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
   598
        by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   599
      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
   600
        by (auto simp: mem_box)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   601
      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
   602
        unfolding p_def by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   603
    }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   604
    ultimately show "\<Union>p = cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   605
      by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   606
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   607
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   608
66154
bc5e6461f759 Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents: 66113
diff changeset
   609
proposition partial_division_extend_interval:
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   610
  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
   611
  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
   612
proof (cases "p = {}")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   613
  case True
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   614
  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
   615
    by (rule elementary_interval)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   616
  then show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   617
    using True that by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   618
next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   619
  case False
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   620
  note p = division_ofD[OF assms(1)]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   621
  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
   622
  proof
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   623
    fix k
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   624
    assume kp: "k \<in> p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   625
    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
   626
      using p(4)[OF kp] by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   627
    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
   628
      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
   629
      by (blast intro: order.trans)+
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   630
    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
   631
      by (rule partial_division_extend_1[OF *])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   632
    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
   633
      unfolding k by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   634
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   635
  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
   636
    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
   637
  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
   638
    apply (rule division_ofI)
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   639
    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
   640
  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
   641
    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
   642
  have "\<exists>d. d division_of \<Inter>((\<lambda>i. \<Union>(q i - {i})) ` p)"
66317
a9bb833ee971 more tidying up
paulson
parents: 66314
diff changeset
   643
    apply (rule elementary_Inter [OF finite_imageI[OF p(1)]])
a9bb833ee971 more tidying up
paulson
parents: 66314
diff changeset
   644
    apply (auto dest: di simp: False elementary_Inter [OF finite_imageI[OF p(1)]])
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   645
    done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   646
  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
   647
  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
   648
  proof -
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   649
    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
   650
    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
   651
    proof (rule te[OF False], clarify)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   652
      fix i
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   653
      assume i: "i \<in> p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   654
      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
   655
        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
   656
    qed
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   657
    { fix K
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   658
      assume K: "K \<in> p"
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   659
      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
   660
      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
   661
        by auto
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   662
      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
   663
      proof (rule *[OF Int_interior_Union_intervals])
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   664
        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
   665
          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
   666
        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
   667
          apply (rule interior_mono)+
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   668
          using K by auto
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   669
      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
   670
    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
   671
      unfolding cbox_eq
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   672
      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
   673
      apply (rule Int_interior_Union_intervals)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   674
      apply (rule p open_interior ballI)+
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   675
      apply simp_all
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   676
      done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   677
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   678
  then show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   679
    by (meson Un_upper2 that)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   680
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   681
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   682
lemma elementary_bounded[dest]:
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   683
  fixes S :: "'a::euclidean_space set"
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   684
  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
   685
  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
   686
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   687
lemma elementary_subset_cbox:
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   688
  "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
   689
  by (meson elementary_bounded bounded_subset_cbox)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   690
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   691
lemma division_union_intervals_exists:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   692
  fixes a b :: "'a::euclidean_space"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   693
  assumes "cbox a b \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   694
  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
   695
proof (cases "cbox c d = {}")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   696
  case True
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
   697
  with assms that show ?thesis by force
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   698
next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   699
  case False
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   700
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   701
  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
   702
    case True
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 (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
   705
  next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   706
    case False
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   707
    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
   708
      unfolding Int_interval by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   709
    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
   710
    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
   711
      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
   712
    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
   713
    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
   714
      apply (rule arg_cong[of _ _ interior])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   715
      using p(8) uv by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   716
    also have "\<dots> = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   717
      unfolding interior_Int
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   718
      apply (rule Int_interior_Union_intervals)
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   719
      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
   720
      apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   721
      done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   722
    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
   723
    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
   724
      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
   725
    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
   726
    proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   727
      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
   728
        by (simp add: assms division_of_self)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   729
      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
   730
        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
   731
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   732
    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
   733
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   734
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   735
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   736
lemma division_of_Union:
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   737
  assumes "finite f"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   738
    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
   739
    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
   740
  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
   741
  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
   742
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   743
lemma elementary_union_interval:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   744
  fixes a b :: "'a::euclidean_space"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   745
  assumes "p division_of \<Union>p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   746
  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
   747
proof (cases "p = {} \<or> cbox a b = {}")
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   748
  case True
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   749
  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
   750
    by (rule elementary_interval)
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   751
  then show thesis
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   752
    using True assms that by auto
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   753
next
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   754
  case False
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   755
  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
   756
    by auto
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   757
  note pdiv = division_ofD[OF assms]
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   758
  show ?thesis
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   759
  proof (cases "interior (cbox a b) = {}")
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   760
    case True
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   761
    show ?thesis
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   762
      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
   763
      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
   764
       apply (metis IntI equals0D pdiv(5))
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   765
      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
   766
  next
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   767
    case False  
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   768
    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
   769
    proof
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   770
      fix K
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   771
      assume kp: "K \<in> p"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   772
      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
   773
      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
   774
        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
   775
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   776
    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
   777
    note q = division_ofD[OF this[rule_format]]
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   778
    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
   779
    show thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   780
    proof (rule that[OF division_ofI])
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   781
      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
   782
        by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   783
      show "finite ?D"
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   784
        using "*" pdiv(1) q(1) by auto
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
   785
      have "\<Union>?D = (\<Union>x\<in>p. \<Union>insert (cbox a b) (q x))"
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   786
        by auto
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
   787
      also have "... = \<Union>{cbox a b \<union> t |t. t \<in> p}"
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
   788
        using q(6) by auto
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
   789
      also have "... = cbox a b \<union> \<Union>p"
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
   790
        using \<open>p \<noteq> {}\<close> by auto
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
   791
      finally show "\<Union>?D = cbox a b \<union> \<Union>p" .
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   792
      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
   793
        using q that by blast+
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   794
      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
   795
        using q(4) that by auto
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   796
    next
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   797
      fix K' K
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   798
      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
   799
      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
   800
        using K by auto
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   801
      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
   802
        using K' by auto
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   803
      show "interior K \<inter> interior K' = {}"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   804
      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
   805
        case True
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   806
        then show ?thesis
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   807
          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
   808
      next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   809
        case False
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   810
        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
   811
          by auto
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   812
        obtain c d where K: "K = cbox c d"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   813
          using q(4) x by blast
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   814
        have "interior K \<inter> interior (cbox a b) = {}"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   815
          using as' K'(2) q(5) x by blast
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   816
        then have "interior K \<subseteq> interior x"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   817
          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
   818
        moreover
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   819
        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
   820
          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
   821
        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
   822
          using as'(2) q(5) x' by blast
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   823
        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
   824
          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
   825
        moreover have "interior x \<inter> interior x' = {}"
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   826
          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
   827
        ultimately show ?thesis
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   828
          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
   829
      qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   830
    qed
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   831
  qed
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   832
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   833
66299
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   834
1b4aa3e3e4e6 partial cleanup of the horrible Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66296
diff changeset
   835
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   836
lemma elementary_unions_intervals:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   837
  assumes fin: "finite f"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   838
    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
   839
  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
   840
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   841
  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
   842
  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
   843
    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
   844
  next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   845
    fix x F
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   846
    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
   847
    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
   848
    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
   849
    have *: "\<Union>F = \<Union>p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   850
      using division_ofD[OF p] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   851
    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
   852
      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
   853
      unfolding Union_insert x * by metis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   854
  qed (insert assms, auto)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   855
  then show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   856
    using that by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   857
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   858
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   859
lemma elementary_union:
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   860
  fixes S T :: "'a::euclidean_space set"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   861
  assumes "ps division_of S" "pt division_of T"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   862
  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
   863
proof -
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   864
  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
   865
    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
   866
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   867
    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
   868
    using assms apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   869
    by (simp add: * that)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   870
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   871
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   872
lemma partial_division_extend:
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   873
  fixes T :: "'a::euclidean_space set"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   874
  assumes "p division_of S"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   875
    and "q division_of T"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   876
    and "S \<subseteq> T"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   877
  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
   878
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   879
  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
   880
  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
   881
    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
   882
  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
   883
    using assms
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   884
    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
   885
  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
   886
  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
   887
    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
   888
    using r1(3,6)
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   889
      apply auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   890
    done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   891
  then obtain r2 where r2: "r2 division_of (\<Union>(r1 - p)) \<inter> (\<Union>q)"
66317
a9bb833ee971 more tidying up
paulson
parents: 66314
diff changeset
   892
    by (metis assms(2) divq(6) elementary_Int)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   893
  {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   894
    fix x
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   895
    assume x: "x \<in> T" "x \<notin> S"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   896
    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
   897
      unfolding r1 using ab
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   898
      by (meson division_contains r1(2) subsetCE)
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   899
    moreover have "R \<notin> p"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   900
    proof
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   901
      assume "R \<in> p"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   902
      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
   903
      then show False using x by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   904
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   905
    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
   906
  }
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   907
  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
   908
    unfolding divp divq using assms(3) by auto
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   909
  have "interior S \<inter> interior (\<Union>(r1-p)) = {}"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   910
  proof (rule Int_interior_Union_intervals)
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   911
    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
   912
      by auto
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   913
    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
   914
    proof -
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   915
      have "interior m \<inter> interior (\<Union>p) = {}"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   916
      proof (rule Int_interior_Union_intervals)
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   917
        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
   918
          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
   919
      qed (use divp in auto)
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   920
      then show "interior S \<inter> interior m = {}"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   921
        unfolding divp by auto
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   922
    qed 
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   923
  qed (use r1 in auto)
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   924
  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
   925
    using interior_subset by auto
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   926
  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
   927
    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
   928
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   929
    apply (rule that[of "p \<union> r2"])
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   930
     apply (auto simp: div Teq)
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   931
    done
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   932
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   933
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   934
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   935
lemma division_split:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   936
  fixes a :: "'a::euclidean_space"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   937
  assumes "p division_of (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   938
    and k: "k\<in>Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   939
  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
   940
      (is "?p1 division_of ?I1")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   941
    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
   942
      (is "?p2 division_of ?I2")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   943
proof (rule_tac[!] division_ofI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   944
  note p = division_ofD[OF assms(1)]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   945
  show "finite ?p1" "finite ?p2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   946
    using p(1) by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   947
  show "\<Union>?p1 = ?I1" "\<Union>?p2 = ?I2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   948
    unfolding p(6)[symmetric] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   949
  {
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   950
    fix K
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   951
    assume "K \<in> ?p1"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   952
    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
   953
      by blast
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   954
    obtain u v where uv: "l = cbox u v"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   955
      using assms(1) l(2) by blast
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   956
    show "K \<subseteq> ?I1"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   957
      using l p(2) uv by force
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   958
    show  "K \<noteq> {}"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   959
      by (simp add: l)
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   960
    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
   961
      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
   962
      apply (subst interval_split[OF k])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   963
      apply (auto intro: order.trans)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   964
      done
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   965
    fix K'
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   966
    assume "K' \<in> ?p1"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   967
    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
   968
      by blast
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   969
    assume "K \<noteq> K'"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   970
    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
   971
      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
   972
  }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   973
  {
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   974
    fix K
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   975
    assume "K \<in> ?p2"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   976
    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
   977
      by blast
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   978
    obtain u v where uv: "l = cbox u v"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   979
      using l(2) p(4) by blast
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   980
    show "K \<subseteq> ?I2"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   981
      using l p(2) uv by force
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   982
    show  "K \<noteq> {}"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   983
      by (simp add: l)
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   984
    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
   985
      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
   986
      apply (subst interval_split[OF k])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   987
      apply (auto intro: order.trans)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   988
      done
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   989
    fix K'
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   990
    assume "K' \<in> ?p2"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   991
    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
   992
      by blast
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   993
    assume "K \<noteq> K'"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
   994
    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
   995
      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
   996
  }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   997
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   998
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
   999
subsection \<open>Tagged (partial) divisions.\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1000
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1001
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
  1002
  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
  1003
    finite s \<and>
66113
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1004
    (\<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
  1005
    (\<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
  1006
      interior K1 \<inter> interior K2 = {})"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1007
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1008
lemma tagged_partial_division_ofD[dest]:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1009
  assumes "s tagged_partial_division_of i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1010
  shows "finite s"
66113
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1011
    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
  1012
    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
  1013
    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
  1014
    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
  1015
      (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
  1016
  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
  1017
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1018
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
  1019
  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
  1020
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1021
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
  1022
  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
  1023
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1024
lemma tagged_division_of:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1025
  "s tagged_division_of i \<longleftrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1026
    finite s \<and>
66113
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1027
    (\<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
  1028
    (\<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
  1029
      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
  1030
    (\<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
  1031
  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
  1032
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1033
lemma tagged_division_ofI:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1034
  assumes "finite s"
66113
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1035
    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
  1036
    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
  1037
    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
  1038
    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
  1039
      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
  1040
    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
  1041
  shows "s tagged_division_of i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1042
  unfolding tagged_division_of
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1043
  using assms  by fastforce
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1044
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1045
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
  1046
  assumes "s tagged_division_of i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1047
  shows "finite s"
66113
571b698659c0 Repaired an inadvertent reordering of the premises of two theorems
paulson <lp15@cam.ac.uk>
parents: 66112
diff changeset
  1048
    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
  1049
    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
  1050
    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
  1051
    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
  1052
      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
  1053
    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
  1054
  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
  1055
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1056
lemma division_of_tagged_division:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1057
  assumes "s tagged_division_of i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1058
  shows "(snd ` s) division_of i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1059
proof (rule division_ofI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1060
  note assm = tagged_division_ofD[OF assms]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1061
  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
  1062
    using assm by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1063
  fix k
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1064
  assume k: "k \<in> snd ` s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1065
  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
  1066
    by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1067
  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
  1068
    using assm by fastforce+
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1069
  fix k'
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1070
  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
  1071
  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
  1072
    by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1073
  then show "interior k \<inter> interior k' = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1074
    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
  1075
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1076
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1077
lemma partial_division_of_tagged_division:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1078
  assumes "s tagged_partial_division_of i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1079
  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
  1080
proof (rule division_ofI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1081
  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
  1082
  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
  1083
    using assm by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1084
  fix k
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1085
  assume k: "k \<in> snd ` s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1086
  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
  1087
    by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1088
  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
  1089
    using assm by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1090
  fix k'
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1091
  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
  1092
  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
  1093
    by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1094
  then show "interior k \<inter> interior k' = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1095
    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
  1096
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1097
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1098
lemma tagged_partial_division_subset:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1099
  assumes "s tagged_partial_division_of i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1100
    and "t \<subseteq> s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1101
  shows "t tagged_partial_division_of i"
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1102
  using assms finite_subset[OF assms(2)]
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1103
  unfolding tagged_partial_division_of_def
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1104
  by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1105
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1106
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
  1107
  by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1108
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1109
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
  1110
  unfolding tagged_division_of by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1111
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1112
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
  1113
  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
  1114
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1115
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
  1116
  unfolding tagged_division_of by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1117
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1118
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
  1119
  by (rule tagged_division_ofI) auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1120
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1121
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
  1122
  unfolding box_real[symmetric]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1123
  by (rule tagged_division_of_self)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1124
66497
18a6478a574c More tidying, and renaming of theorems
paulson <lp15@cam.ac.uk>
parents: 66365
diff changeset
  1125
lemma tagged_division_Un:
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1126
  assumes "p1 tagged_division_of s1"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1127
    and "p2 tagged_division_of s2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1128
    and "interior s1 \<inter> interior s2 = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1129
  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
  1130
proof (rule tagged_division_ofI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1131
  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
  1132
  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
  1133
  show "finite (p1 \<union> p2)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1134
    using p1(1) p2(1) by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1135
  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
  1136
    using p1(6) p2(6) by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1137
  fix x k
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1138
  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
  1139
  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
  1140
    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
  1141
  show "k \<subseteq> s1 \<union> s2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1142
    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
  1143
  fix x' k'
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1144
  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
  1145
  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
  1146
    using assms(3) interior_mono by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1147
  show "interior k \<inter> interior k' = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1148
    apply (cases "(x, k) \<in> p1")
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1149
    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
  1150
    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
  1151
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1152
66497
18a6478a574c More tidying, and renaming of theorems
paulson <lp15@cam.ac.uk>
parents: 66365
diff changeset
  1153
lemma tagged_division_Union:
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1154
  assumes "finite I"
66497
18a6478a574c More tidying, and renaming of theorems
paulson <lp15@cam.ac.uk>
parents: 66365
diff changeset
  1155
    and tag: "\<And>i. i\<in>I \<Longrightarrow> pfn i tagged_division_of i"
18a6478a574c More tidying, and renaming of theorems
paulson <lp15@cam.ac.uk>
parents: 66365
diff changeset
  1156
    and disj: "\<And>i1 i2. \<lbrakk>i1 \<in> I; i2 \<in> I; i1 \<noteq> i2\<rbrakk> \<Longrightarrow> interior(i1) \<inter> interior(i2) = {}"
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1157
  shows "\<Union>(pfn ` I) tagged_division_of (\<Union>I)"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1158
proof (rule tagged_division_ofI)
66497
18a6478a574c More tidying, and renaming of theorems
paulson <lp15@cam.ac.uk>
parents: 66365
diff changeset
  1159
  note assm = tagged_division_ofD[OF tag]
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1160
  show "finite (\<Union>(pfn ` I))"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1161
    using assms by auto
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1162
  have "\<Union>{k. \<exists>x. (x, k) \<in> \<Union>(pfn ` I)} = \<Union>((\<lambda>i. \<Union>{k. \<exists>x. (x, k) \<in> pfn i}) ` I)"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1163
    by blast
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1164
  also have "\<dots> = \<Union>I"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1165
    using assm(6) by auto
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1166
  finally show "\<Union>{k. \<exists>x. (x, k) \<in> \<Union>(pfn ` I)} = \<Union>I" .
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1167
  fix x k
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1168
  assume xk: "(x, k) \<in> \<Union>(pfn ` I)"
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1169
  then obtain i where i: "i \<in> I" "(x, k) \<in> pfn i"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1170
    by auto
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1171
  show "x \<in> k" "\<exists>a b. k = cbox a b" "k \<subseteq> \<Union>I"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1172
    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
  1173
  fix x' k'
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1174
  assume xk': "(x', k') \<in> \<Union>(pfn ` I)" "(x, k) \<noteq> (x', k')"
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1175
  then obtain i' where i': "i' \<in> I" "(x', k') \<in> pfn i'"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1176
    by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1177
  have *: "\<And>a b. i \<noteq> i' \<Longrightarrow> a \<subseteq> i \<Longrightarrow> b \<subseteq> i' \<Longrightarrow> interior a \<inter> interior b = {}"
66497
18a6478a574c More tidying, and renaming of theorems
paulson <lp15@cam.ac.uk>
parents: 66365
diff changeset
  1178
    using i(1) i'(1) disj interior_mono by blast
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1179
  show "interior k \<inter> interior k' = {}"
66497
18a6478a574c More tidying, and renaming of theorems
paulson <lp15@cam.ac.uk>
parents: 66365
diff changeset
  1180
  proof (cases "i = i'")
18a6478a574c More tidying, and renaming of theorems
paulson <lp15@cam.ac.uk>
parents: 66365
diff changeset
  1181
    case True then show ?thesis 
18a6478a574c More tidying, and renaming of theorems
paulson <lp15@cam.ac.uk>
parents: 66365
diff changeset
  1182
      using assm(5) i' i xk'(2) by blast
18a6478a574c More tidying, and renaming of theorems
paulson <lp15@cam.ac.uk>
parents: 66365
diff changeset
  1183
  next
18a6478a574c More tidying, and renaming of theorems
paulson <lp15@cam.ac.uk>
parents: 66365
diff changeset
  1184
    case False then show ?thesis 
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1185
    using "*" assm(3) i' i by auto
66497
18a6478a574c More tidying, and renaming of theorems
paulson <lp15@cam.ac.uk>
parents: 66365
diff changeset
  1186
  qed
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1187
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1188
66497
18a6478a574c More tidying, and renaming of theorems
paulson <lp15@cam.ac.uk>
parents: 66365
diff changeset
  1189
lemma tagged_partial_division_of_Union_self:
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1190
  assumes "p tagged_partial_division_of s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1191
  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
  1192
  apply (rule tagged_division_ofI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1193
  using tagged_partial_division_ofD[OF assms]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1194
  apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1195
  done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1196
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1197
lemma tagged_division_of_union_self:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1198
  assumes "p tagged_division_of s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1199
  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
  1200
  apply (rule tagged_division_ofI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1201
  using tagged_division_ofD[OF assms]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1202
  apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1203
  done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1204
66498
97fc319d6089 Merge (non-trivial)
paulson <lp15@cam.ac.uk>
parents: 66492 66497
diff changeset
  1205
lemma tagged_division_Un_interval:
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1206
  fixes a :: "'a::euclidean_space"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1207
  assumes "p1 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<le> (c::real)})"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1208
    and "p2 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1209
    and k: "k \<in> Basis"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1210
  shows "(p1 \<union> p2) tagged_division_of (cbox a b)"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1211
proof -
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1212
  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})"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1213
    by auto
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1214
  show ?thesis
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1215
    apply (subst *)
66498
97fc319d6089 Merge (non-trivial)
paulson <lp15@cam.ac.uk>
parents: 66492 66497
diff changeset
  1216
    apply (rule tagged_division_Un[OF assms(1-2)])
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1217
    unfolding interval_split[OF k] interior_cbox
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1218
    using k
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1219
    apply (auto simp add: box_def elim!: ballE[where x=k])
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1220
    done
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1221
qed
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1222
66498
97fc319d6089 Merge (non-trivial)
paulson <lp15@cam.ac.uk>
parents: 66492 66497
diff changeset
  1223
lemma tagged_division_Un_interval_real:
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1224
  fixes a :: real
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1225
  assumes "p1 tagged_division_of ({a .. b} \<inter> {x. x\<bullet>k \<le> (c::real)})"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1226
    and "p2 tagged_division_of ({a .. b} \<inter> {x. x\<bullet>k \<ge> c})"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1227
    and k: "k \<in> Basis"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1228
  shows "(p1 \<union> p2) tagged_division_of {a .. b}"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1229
  using assms
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1230
  unfolding box_real[symmetric]
66498
97fc319d6089 Merge (non-trivial)
paulson <lp15@cam.ac.uk>
parents: 66492 66497
diff changeset
  1231
  by (rule tagged_division_Un_interval)
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1232
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1233
lemma tagged_division_split_left_inj:
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1234
  assumes d: "d tagged_division_of i"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1235
  and tags: "(x1, K1) \<in> d" "(x2, K2) \<in> d"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1236
  and "K1 \<noteq> K2"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1237
  and eq: "K1 \<inter> {x. x \<bullet> k \<le> c} = K2 \<inter> {x. x \<bullet> k \<le> c}"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1238
    shows "interior (K1 \<inter> {x. x\<bullet>k \<le> c}) = {}"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1239
proof -
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1240
  have "interior (K1 \<inter> K2) = {} \<or> (x2, K2) = (x1, K1)"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1241
    using tags d by (metis (no_types) interior_Int tagged_division_ofD(5))
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1242
  then show ?thesis
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1243
    using eq \<open>K1 \<noteq> K2\<close> by (metis (no_types) inf_assoc inf_bot_left inf_left_idem interior_Int old.prod.inject)
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1244
qed
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1245
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1246
lemma tagged_division_split_right_inj:
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1247
  assumes d: "d tagged_division_of i"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1248
  and tags: "(x1, K1) \<in> d" "(x2, K2) \<in> d"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1249
  and "K1 \<noteq> K2"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1250
  and eq: "K1 \<inter> {x. x\<bullet>k \<ge> c} = K2 \<inter> {x. x\<bullet>k \<ge> c}"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1251
    shows "interior (K1 \<inter> {x. x\<bullet>k \<ge> c}) = {}"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1252
proof -
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1253
  have "interior (K1 \<inter> K2) = {} \<or> (x2, K2) = (x1, K1)"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1254
    using tags d by (metis (no_types) interior_Int tagged_division_ofD(5))
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1255
  then show ?thesis
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1256
    using eq \<open>K1 \<noteq> K2\<close> by (metis (no_types) inf_assoc inf_bot_left inf_left_idem interior_Int old.prod.inject)
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1257
qed
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1258
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1259
lemma (in comm_monoid_set) over_tagged_division_lemma:
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1260
  assumes "p tagged_division_of i"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1261
    and "\<And>u v. box u v = {} \<Longrightarrow> d (cbox u v) = \<^bold>1"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1262
  shows "F (\<lambda>(_, k). d k) p = F d (snd ` p)"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1263
proof -
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1264
  have *: "(\<lambda>(_ ,k). d k) = d \<circ> snd"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1265
    by (simp add: fun_eq_iff)
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1266
  note assm = tagged_division_ofD[OF assms(1)]
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1267
  show ?thesis
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1268
    unfolding *
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1269
  proof (rule reindex_nontrivial[symmetric])
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1270
    show "finite p"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1271
      using assm by auto
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1272
    fix x y
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1273
    assume "x\<in>p" "y\<in>p" "x\<noteq>y" "snd x = snd y"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1274
    obtain a b where ab: "snd x = cbox a b"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1275
      using assm(4)[of "fst x" "snd x"] \<open>x\<in>p\<close> by auto
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1276
    have "(fst x, snd y) \<in> p" "(fst x, snd y) \<noteq> y"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1277
      using \<open>x \<in> p\<close> \<open>x \<noteq> y\<close> \<open>snd x = snd y\<close> [symmetric] by auto
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1278
    with \<open>x\<in>p\<close> \<open>y\<in>p\<close> have "interior (snd x) \<inter> interior (snd y) = {}"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1279
      by (intro assm(5)[of "fst x" _ "fst y"]) auto
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1280
    then have "box a b = {}"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1281
      unfolding \<open>snd x = snd y\<close>[symmetric] ab by auto
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1282
    then have "d (cbox a b) = \<^bold>1"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1283
      using assm(2)[of "fst x" "snd x"] \<open>x\<in>p\<close> ab[symmetric] by (intro assms(2)) auto
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1284
    then show "d (snd x) = \<^bold>1"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1285
      unfolding ab by auto
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1286
  qed
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1287
qed
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1288
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1289
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1290
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
  1291
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1292
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
  1293
  \<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
  1294
  @{typ bool}.\<close>
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
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
  1297
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1298
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
  1299
where
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1300
  "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
  1301
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1302
lemma lift_option_simps[simp]:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1303
  "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
  1304
  "lift_option f None b' = None"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1305
  "lift_option f a' None = None"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1306
  by (auto simp: lift_option_def)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1307
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1308
lemma comm_monoid_lift_option:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1309
  assumes "comm_monoid f z"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1310
  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
  1311
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1312
  from assms interpret comm_monoid f z .
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1313
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1314
    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
  1315
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1316
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1317
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
  1318
  by standard auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1319
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1320
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
  1321
  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
  1322
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1323
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1324
paragraph \<open>Misc\<close>
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1325
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1326
lemma interval_real_split:
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1327
  "{a .. b::real} \<inter> {x. x \<le> c} = {a .. min b c}"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1328
  "{a .. b} \<inter> {x. c \<le> x} = {max a c .. b}"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1329
  apply (metis Int_atLeastAtMostL1 atMost_def)
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1330
  apply (metis Int_atLeastAtMostL2 atLeast_def)
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1331
  done
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1332
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1333
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)"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1334
  by (meson zero_less_one)
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1335
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1336
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1337
paragraph \<open>Division points\<close>
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1338
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1339
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
  1340
   {(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
  1341
     (\<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
  1342
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1343
lemma division_points_finite:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1344
  fixes i :: "'a::euclidean_space set"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1345
  assumes "d division_of i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1346
  shows "finite (division_points i d)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1347
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1348
  note assm = division_ofD[OF assms]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1349
  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
  1350
    (\<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
  1351
  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
  1352
    unfolding division_points_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1353
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1354
    unfolding * using assm by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1355
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1356
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1357
lemma division_points_subset:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1358
  fixes a :: "'a::euclidean_space"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1359
  assumes "d division_of (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1360
    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
  1361
    and k: "k \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1362
  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
  1363
      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
  1364
    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
  1365
      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
  1366
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1367
  note assm = division_ofD[OF assms(1)]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1368
  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
  1369
    "\<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
  1370
    "\<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
  1371
    "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
  1372
    using assms using less_imp_le by auto
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1373
   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
  1374
     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
  1375
        "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
  1376
        "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
  1377
        "x \<in> Basis" for i l x y
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1378
  proof -
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1379
    obtain u v where l: "l = cbox u v"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1380
      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
  1381
    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
  1382
      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
  1383
    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
  1384
      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
  1385
    show ?thesis
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1386
      apply (rule bexI[OF _ \<open>l \<in> d\<close>])
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1387
      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
  1388
      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
  1389
      apply (auto split: if_split_asm)
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1390
      done
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1391
  qed
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1392
  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
  1393
    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
  1394
  ultimately show ?t1 
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1395
    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
  1396
    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
  1397
    unfolding * by force
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1398
  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
  1399
    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
  1400
  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
  1401
                       interval_upperbound i \<bullet> x = y"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1402
    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
  1403
      "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
  1404
      "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
  1405
      "x \<in> Basis" for x y i l
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1406
  proof -
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1407
    obtain u v where l: "l = cbox u v"
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1408
      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
  1409
    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
  1410
      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
  1411
    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
  1412
      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
  1413
    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
  1414
      apply (rule bexI[OF _ \<open>l \<in> d\<close>])
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1415
      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
  1416
      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
  1417
      apply (auto split: if_split_asm)
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1418
      done
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1419
  qed
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1420
  ultimately show ?t2
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1421
    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
  1422
    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
  1423
    unfolding *
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1424
    by force
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1425
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1426
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1427
lemma division_points_psubset:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1428
  fixes a :: "'a::euclidean_space"
66314
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  1429
  assumes d: "d division_of (cbox a b)"
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  1430
      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
  1431
      and "l \<in> d"
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1432
      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
  1433
      and k: "k \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1434
  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
  1435
         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
  1436
    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
  1437
         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
  1438
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1439
  have ab: "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i"
66314
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  1440
    using altb by (auto intro!:less_imp_le)
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  1441
  obtain u v where l: "l = cbox u v"
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  1442
    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
  1443
  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
  1444
    apply (metis assms(5) box_ne_empty(1) cbox_division_memE d l)
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  1445
    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
  1446
  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
  1447
          "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
  1448
    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
  1449
    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
  1450
    by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1451
  have "\<exists>x. x \<in> ?D - ?D1"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1452
    using assms(3-)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1453
    unfolding division_points_def interval_bounds[OF ab]
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1454
    by (force simp add: *)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1455
  moreover have "?D1 \<subseteq> ?D"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1456
    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
  1457
  ultimately show "?D1 \<subset> ?D"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1458
    by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1459
  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
  1460
    "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
  1461
    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
  1462
    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
  1463
    by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1464
  have "\<exists>x. x \<in> ?D - ?D2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1465
    using assms(3-)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1466
    unfolding division_points_def interval_bounds[OF ab]
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1467
    by (force simp add: *)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1468
  moreover have "?D2 \<subseteq> ?D"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1469
    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
  1470
  ultimately show "?D2 \<subset> ?D"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1471
    by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1472
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1473
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1474
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
  1475
  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
  1476
  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
  1477
    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
  1478
    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
  1479
  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
  1480
proof -
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1481
  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
  1482
    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
  1483
  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
  1484
    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
  1485
  ultimately show ?thesis
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1486
    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
  1487
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1488
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1489
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
  1490
  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
  1491
  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
  1492
    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
  1493
    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
  1494
  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
  1495
proof -
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1496
  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
  1497
    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
  1498
  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
  1499
    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
  1500
  ultimately show ?thesis
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1501
    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
  1502
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1503
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1504
lemma interval_doublesplit:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1505
  fixes a :: "'a::euclidean_space"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1506
  assumes "k \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1507
  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
  1508
    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
  1509
     (\<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
  1510
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1511
  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
  1512
    by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1513
  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
  1514
    by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1515
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1516
    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
  1517
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1518
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1519
lemma division_doublesplit:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1520
  fixes a :: "'a::euclidean_space"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1521
  assumes "p division_of (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1522
    and k: "k \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1523
  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
  1524
         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
  1525
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1526
  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
  1527
    by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1528
  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
  1529
    by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1530
  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
  1531
  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
  1532
  then show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1533
    apply (rule **)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1534
    subgoal
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1535
      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
  1536
      apply (rule equalityI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1537
      apply blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1538
      apply clarsimp
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1539
      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
  1540
      apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1541
      done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1542
    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
  1543
qed
66306
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1544
              
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1545
paragraph \<open>Operative\<close>
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1546
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1547
locale operative = comm_monoid_set +
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1548
  fixes g :: "'b::euclidean_space set \<Rightarrow> 'a"
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1549
  assumes box_empty_imp: "\<And>a b. box a b = {} \<Longrightarrow> g (cbox a b) = \<^bold>1"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1550
    and Basis_imp: "\<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})"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1551
begin
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1552
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1553
lemma empty [simp]:
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1554
  "g {} = \<^bold>1"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1555
proof -
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1556
  have *: "cbox One (-One) = ({}::'b set)"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1557
    by (auto simp: box_eq_empty inner_sum_left inner_Basis sum.If_cases ex_in_conv)
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1558
  moreover have "box One (-One) = ({}::'b set)"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1559
    using box_subset_cbox[of One "-One"] by (auto simp: *)
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1560
  ultimately show ?thesis
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1561
    using box_empty_imp [of One "-One"] by simp
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1562
qed
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1563
       
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1564
lemma division:
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1565
  "F g d = g (cbox a b)" if "d division_of (cbox a b)"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1566
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1567
  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
  1568
  then show ?thesis
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1569
  using that proof (induction C arbitrary: a b d rule: less_induct)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1570
    case (less a b d)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1571
    show ?case
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1572
    proof cases
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1573
      assume "box a b = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1574
      { fix k assume "k\<in>d"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1575
        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
  1576
          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
  1577
        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
  1578
          by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1579
        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
  1580
          unfolding subset_box by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1581
        then have "g k = \<^bold>1"
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1582
          using box_empty_imp [of a' b'] k by (simp add: \<open>box a b = {}\<close>) }
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1583
      then show "box a b = {} \<Longrightarrow> F g d = g (cbox a b)"
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1584
        by (auto intro!: neutral simp: box_empty_imp)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1585
    next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1586
      assume "box a b \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1587
      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
  1588
        by (auto simp: box_ne_empty)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1589
      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
  1590
      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
  1591
        case True
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1592
        { fix u v and j :: 'b
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1593
          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
  1594
          then have "cbox u v \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1595
            using less.prems by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1596
          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
  1597
            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
  1598
          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
  1599
            using as j by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1600
          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
  1601
               "(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
  1602
          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
  1603
          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
  1604
          moreover
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1605
          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
  1606
            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
  1607
            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
  1608
            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
  1609
          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
  1610
            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
  1611
        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
  1612
          (\<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
  1613
          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
  1614
        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
  1615
          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
  1616
        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
  1617
        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
  1618
        obtain u v where uv: "i = cbox u v" 
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1619
                     "\<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
  1620
                                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
  1621
                                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
  1622
          using d' i(1) by auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1623
        have "cbox a b \<in> d"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1624
        proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1625
          have "u = a" "v = b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1626
            unfolding euclidean_eq_iff[where 'a='b]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1627
          proof safe
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1628
            fix j :: 'b
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1629
            assume j: "j \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1630
            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
  1631
            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
  1632
              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
  1633
          qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1634
          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
  1635
          then show ?thesis using i by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1636
        qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1637
        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
  1638
          by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1639
        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
  1640
        proof (intro neutral ballI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1641
          fix x
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1642
          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
  1643
          then have "x\<in>d"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1644
            by auto note d'[rule_format,OF this]
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  1645
          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
  1646
                      "\<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
  1647
                                 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
  1648
                                 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
  1649
            by blast
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1650
          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
  1651
            using x[unfolded uv] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1652
          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
  1653
            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
  1654
          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
  1655
            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
  1656
          then have "box u v = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1657
            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
  1658
          then show "g x = \<^bold>1"
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1659
            unfolding uv(1) by (rule box_empty_imp)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1660
        qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1661
        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
  1662
          using division_ofD[OF less.prems]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1663
          apply (subst deq)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1664
          apply (subst insert)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1665
          apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1666
          done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1667
      next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1668
        case False
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1669
        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
  1670
          by auto
66306
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1671
        then obtain k c 
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1672
          where "k \<in> Basis" "interval_lowerbound (cbox a b) \<bullet> k < c"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1673
                "c < interval_upperbound (cbox a b) \<bullet> k"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1674
                "\<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
  1675
          unfolding division_points_def by auto
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1676
        then obtain j where "a \<bullet> k < c" "c < b \<bullet> k" 
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1677
              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
  1678
          by (metis division_of_trivial empty_iff interval_bounds' less.prems)
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1679
        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
  1680
        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
  1681
        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
  1682
        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
  1683
        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
  1684
        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
  1685
              \<subset> division_points (cbox a b) d"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1686
          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
  1687
        with division_points_finite[OF \<open>d division_of cbox a b\<close>] 
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1688
        have "card
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1689
         (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
  1690
           < card (division_points (cbox a b) d)"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1691
          by (rule psubset_card_mono)
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1692
        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
  1693
              \<subset> division_points (cbox a b) d"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1694
          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
  1695
        with division_points_finite[OF \<open>d division_of cbox a b\<close>] 
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1696
        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
  1697
              < card (division_points (cbox a b) d)"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1698
          by (rule psubset_card_mono)
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1699
        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
  1700
          unfolding interval_split[OF \<open>k \<in> Basis\<close>]
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1701
          apply (rule_tac[!] less.hyps)
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1702
          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
  1703
          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
  1704
        have fxk_le: "g (l \<inter> ?lec) = \<^bold>1" 
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1705
          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
  1706
        proof -
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1707
          obtain u v where leq: "l = cbox u v"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1708
            using \<open>l \<in> d\<close> less.prems by auto
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1709
          have "interior (cbox u v \<inter> ?lec) = {}"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1710
            using that division_split_left_inj leq less.prems by blast
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1711
          then show ?thesis
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1712
            unfolding leq interval_split [OF \<open>k \<in> Basis\<close>]
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1713
            by (auto intro: box_empty_imp)
66306
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1714
        qed
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1715
        have fxk_ge: "g (l \<inter> {x. x \<bullet> k \<ge> c}) = \<^bold>1"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1716
          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
  1717
        proof -
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1718
          obtain u v where leq: "l = cbox u v"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1719
            using \<open>l \<in> d\<close> less.prems by auto
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1720
          have "interior (cbox u v \<inter> ?gec) = {}"
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1721
            using that division_split_right_inj leq less.prems by blast
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1722
          then show ?thesis
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1723
            unfolding leq interval_split[OF \<open>k \<in> Basis\<close>]
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1724
            by (auto intro: box_empty_imp)
66306
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1725
        qed
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1726
        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
  1727
          using d1_def by auto
66306
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1728
        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
  1729
          using d2_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1730
        have "g (cbox a b) = F g d1 \<^bold>* F g d2" (is "_ = ?prev")
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1731
          unfolding * using \<open>k \<in> Basis\<close>
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1732
          by (auto dest: Basis_imp)
66306
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1733
        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
  1734
          unfolding d1_alt using division_of_finite[OF less.prems] fxk_le
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1735
          by (subst reindex_nontrivial) (auto intro!: mono_neutral_cong_left)
66306
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1736
        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
  1737
          unfolding d2_alt using division_of_finite[OF less.prems] fxk_ge
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1738
          by (subst reindex_nontrivial) (auto intro!: mono_neutral_cong_left)
66306
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1739
        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
  1740
          unfolding forall_in_division[OF \<open>d division_of cbox a b\<close>]
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1741
          using \<open>k \<in> Basis\<close>
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1742
          by (auto dest: Basis_imp)
66306
13b051ebc6c5 fixed another horrible proof
paulson
parents: 66300
diff changeset
  1743
        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
  1744
          using * by (simp add: distrib)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1745
        finally show ?thesis by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1746
      qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1747
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1748
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1749
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1750
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1751
lemma tagged_division:
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1752
  assumes "d tagged_division_of (cbox a b)"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1753
  shows "F (\<lambda>(_, l). g l) d = g (cbox a b)"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1754
proof -
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1755
  have "F (\<lambda>(_, k). g k) d = F g (snd ` d)"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1756
    using assms box_empty_imp by (rule over_tagged_division_lemma)
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1757
  then show ?thesis
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1758
    unfolding assms [THEN division_of_tagged_division, THEN division] .
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1759
  qed
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1760
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1761
end
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1762
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1763
locale operative_real = comm_monoid_set +
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1764
  fixes g :: "real set \<Rightarrow> 'a"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1765
  assumes neutral: "b \<le> a \<Longrightarrow> g {a..b} = \<^bold>1"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1766
  assumes coalesce_less: "a < c \<Longrightarrow> c < b \<Longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1767
begin
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1768
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1769
sublocale operative where g = g
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1770
  rewrites "box = (greaterThanLessThan :: real \<Rightarrow> _)"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1771
    and "cbox = (atLeastAtMost :: real \<Rightarrow> _)"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1772
    and "\<And>x::real. x \<in> Basis \<longleftrightarrow> x = 1"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1773
proof -
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1774
  show "operative f z g"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1775
  proof
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1776
    show "g (cbox a b) = \<^bold>1" if "box a b = {}" for a b
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1777
      using that by (simp add: neutral)
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1778
    show "g (cbox a b) = g (cbox a b \<inter> {x. x \<bullet> k \<le> c}) \<^bold>* g (cbox a b \<inter> {x. c \<le> x \<bullet> k})"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1779
      if "k \<in> Basis" for a b c k
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1780
    proof -
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1781
      from that have [simp]: "k = 1"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1782
        by simp
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1783
      from neutral [of 0 1] neutral [of a a for a] coalesce_less
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1784
  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
  1785
    "\<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
  1786
    by auto
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1787
      have "g {a..b} = g {a..min b c} \<^bold>* g {max a c..b}"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1788
    by (auto simp: min_def max_def le_less)
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1789
      then show "g (cbox a b) = g (cbox a b \<inter> {x. x \<bullet> k \<le> c}) \<^bold>* g (cbox a b \<inter> {x. c \<le> x \<bullet> k})"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1790
        by (simp add: atMost_def [symmetric] atLeast_def [symmetric])
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1791
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1792
  qed
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1793
  show "box = (greaterThanLessThan :: real \<Rightarrow> _)"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1794
    and "cbox = (atLeastAtMost :: real \<Rightarrow> _)"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1795
    and "\<And>x::real. x \<in> Basis \<longleftrightarrow> x = 1"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1796
    by (simp_all add: fun_eq_iff)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1797
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1798
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1799
lemma coalesce_less_eq:
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1800
  "g {a..c} \<^bold>* g {c..b} = g {a..b}" if "a \<le> c" "c \<le> b"
66498
97fc319d6089 Merge (non-trivial)
paulson <lp15@cam.ac.uk>
parents: 66492 66497
diff changeset
  1801
  proof (cases "c = a \<or> c = b")
97fc319d6089 Merge (non-trivial)
paulson <lp15@cam.ac.uk>
parents: 66492 66497
diff changeset
  1802
    case False
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1803
  with that have "a < c" "c < b"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1804
    by auto
66498
97fc319d6089 Merge (non-trivial)
paulson <lp15@cam.ac.uk>
parents: 66492 66497
diff changeset
  1805
    then show ?thesis
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1806
    by (rule coalesce_less)
66498
97fc319d6089 Merge (non-trivial)
paulson <lp15@cam.ac.uk>
parents: 66492 66497
diff changeset
  1807
  next
97fc319d6089 Merge (non-trivial)
paulson <lp15@cam.ac.uk>
parents: 66492 66497
diff changeset
  1808
    case True
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1809
  with that box_empty_imp [of a a] box_empty_imp [of b b] show ?thesis
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1810
    by safe simp_all
66498
97fc319d6089 Merge (non-trivial)
paulson <lp15@cam.ac.uk>
parents: 66492 66497
diff changeset
  1811
    qed
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1812
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1813
end
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1814
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1815
lemma operative_realI:
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1816
  "operative_real f z g" if "operative f z g"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1817
proof -
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1818
  interpret operative f z g
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1819
    using that .
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1820
  show ?thesis
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1821
  proof
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1822
    show "g {a..b} = z" if "b \<le> a" for a b
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1823
      using that box_empty_imp by simp
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1824
    show "f (g {a..c}) (g {c..b}) = g {a..b}" if "a < c" "c < b" for a b c
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1825
      using that
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1826
    using Basis_imp [of 1 a b c]
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1827
      by (simp_all add: atMost_def [symmetric] atLeast_def [symmetric] max_def min_def)
66498
97fc319d6089 Merge (non-trivial)
paulson <lp15@cam.ac.uk>
parents: 66492 66497
diff changeset
  1828
qed
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1829
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1830
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1831
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1832
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
  1833
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1834
lemma additive_tagged_division_1:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1835
  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
  1836
  assumes "a \<le> b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1837
    and "p tagged_division_of {a..b}"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
  1838
  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
  1839
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1840
  let ?f = "(\<lambda>k::(real) set. if k = {} then 0 else f(interval_upperbound k) - f(interval_lowerbound k))"
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1841
  interpret operative_real plus 0 ?f
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1842
    rewrites "comm_monoid_set.F op + 0 = sum"
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1843
    by standard[1] (auto simp add: sum_def)
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1844
  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
  1845
    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
  1846
  have **: "cbox a b \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1847
    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
  1848
  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
  1849
    proof -
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1850
      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
  1851
        using assms by auto
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1852
      then show ?thesis
66492
d7206afe2d28 dedicated local for "operative" avoids namespace pollution
haftmann
parents: 66365
diff changeset
  1853
        using p_td assms by (simp add: tagged_division) 
66112
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1854
    qed 
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1855
  then show ?thesis
0e640e04fc56 New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents: 64911
diff changeset
  1856
    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
  1857
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1858
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1859
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1860
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
  1861
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1862
definition fine  (infixr "fine" 46)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1863
  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
  1864
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1865
lemma fineI:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1866
  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
  1867
  shows "d fine s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1868
  using assms unfolding fine_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1869
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1870
lemma fineD[dest]:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1871
  assumes "d fine s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1872
  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
  1873
  using assms unfolding fine_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1874
66192
e5b84854baa4 A few renamings and several tidied-up proofs
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
  1875
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
  1876
  unfolding fine_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1877
66192
e5b84854baa4 A few renamings and several tidied-up proofs
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
  1878
lemma fine_Inter:
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1879
 "(\<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
  1880
  unfolding fine_def by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1881
66192
e5b84854baa4 A few renamings and several tidied-up proofs
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
  1882
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
  1883
  unfolding fine_def by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1884
66192
e5b84854baa4 A few renamings and several tidied-up proofs
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
  1885
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
  1886
  unfolding fine_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1887
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1888
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
  1889
  unfolding fine_def by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1890
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1891
subsection \<open>Some basic combining lemmas.\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1892
66192
e5b84854baa4 A few renamings and several tidied-up proofs
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
  1893
lemma tagged_division_Union_exists:
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1894
  assumes "finite I"
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1895
    and "\<forall>i\<in>I. \<exists>p. p tagged_division_of i \<and> d fine p"
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1896
    and "\<forall>i1\<in>I. \<forall>i2\<in>I. i1 \<noteq> i2 \<longrightarrow> interior i1 \<inter> interior i2 = {}"
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1897
    and "\<Union>I = i"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1898
   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
  1899
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1900
  obtain pfn where pfn:
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1901
    "\<And>x. x \<in> I \<Longrightarrow> pfn x tagged_division_of x"
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1902
    "\<And>x. x \<in> I \<Longrightarrow> d fine pfn x"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1903
    using bchoice[OF assms(2)] by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1904
  show thesis
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1905
    apply (rule_tac p="\<Union>(pfn ` I)" in that)
66497
18a6478a574c More tidying, and renaming of theorems
paulson <lp15@cam.ac.uk>
parents: 66365
diff changeset
  1906
    using assms(1) assms(3) assms(4) pfn(1) tagged_division_Union apply force
66192
e5b84854baa4 A few renamings and several tidied-up proofs
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
  1907
    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
  1908
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1909
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
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
  1912
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1913
lemma division_of_closed:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1914
  fixes i :: "'n::euclidean_space set"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1915
  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
  1916
  unfolding division_of_def by fastforce
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1917
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1918
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
  1919
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1920
lemma interval_bisection_step:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1921
  fixes type :: "'a::euclidean_space"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1922
  assumes "P {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1923
    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
  1924
    and "\<not> P (cbox a (b::'a))"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1925
  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
  1926
    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
  1927
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1928
  have "cbox a b \<noteq> {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1929
    using assms(1,3) by metis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1930
  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
  1931
    by (force simp: mem_box)
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1932
  have UN_cases: "\<lbrakk>finite f;
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1933
           \<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
  1934
           \<And>s. s\<in>f \<Longrightarrow> \<exists>a b. s = cbox a b;
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1935
           \<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)" for f
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1936
  proof (induct f rule: finite_induct)
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1937
    case empty
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1938
    show ?case
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1939
      using assms(1) by auto
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1940
  next
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1941
    case (insert x f)
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1942
    show ?case
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1943
      unfolding Union_insert
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1944
      apply (rule assms(2)[rule_format])
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1945
      using Int_interior_Union_intervals [of f "interior x"]
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1946
      by (metis (no_types, lifting) insert insert_iff open_interior)
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  1947
  qed
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1948
  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
  1949
    (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
  1950
  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
  1951
  {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1952
    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
  1953
    then show thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1954
      unfolding atomize_not not_all
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1955
      by (blast intro: that)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1956
  }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1957
  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
  1958
  have "P (\<Union>?A)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1959
  proof (rule UN_cases)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1960
    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
  1961
      (\<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
  1962
    have "?A \<subseteq> ?B"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1963
    proof
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1964
      fix x
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1965
      assume "x \<in> ?A"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1966
      then obtain c d
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1967
        where x:  "x = cbox c d"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1968
                  "\<And>i. i \<in> Basis \<Longrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1969
                        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
  1970
                        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
  1971
      show "x \<in> ?B"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1972
        unfolding image_iff x
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1973
        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
  1974
        apply (rule arg_cong2 [where f = cbox])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1975
        using x(2) ab
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1976
        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
  1977
        by fastforce
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1978
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1979
    then show "finite ?A"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1980
      by (rule finite_subset) auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1981
  next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1982
    fix s
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1983
    assume "s \<in> ?A"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1984
    then obtain c d
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1985
      where s: "s = cbox c d"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1986
               "\<And>i. i \<in> Basis \<Longrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1987
                     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
  1988
                     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
  1989
      by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1990
    show "P s"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1991
      unfolding s
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1992
      apply (rule as[rule_format])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1993
      using ab s(2) by force
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1994
    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
  1995
      unfolding s by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1996
    fix t
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1997
    assume "t \<in> ?A"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1998
    then obtain e f where t:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  1999
      "t = cbox e f"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2000
      "\<And>i. i \<in> Basis \<Longrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2001
        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
  2002
        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
  2003
      by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2004
    assume "s \<noteq> t"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2005
    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
  2006
      unfolding s t by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2007
    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
  2008
      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
  2009
    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
  2010
      using s(2) t(2) apply fastforce
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2011
      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
  2012
    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
  2013
      by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2014
    show "interior s \<inter> interior t = {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2015
      unfolding s t interior_cbox
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2016
    proof (rule *)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2017
      fix x
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2018
      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
  2019
      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
  2020
        unfolding mem_box using i'
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2021
        by force+
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2022
      show False  using s(2)[OF i']
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2023
      proof safe
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2024
        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
  2025
        show False
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2026
          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
  2027
      next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2028
        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
  2029
        show False
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2030
          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
  2031
      qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2032
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2033
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2034
  also have "\<Union>?A = cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2035
  proof (rule set_eqI,rule)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2036
    fix x
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2037
    assume "x \<in> \<Union>?A"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2038
    then obtain c d where x:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2039
      "x \<in> cbox c d"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2040
      "\<And>i. i \<in> Basis \<Longrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2041
        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
  2042
        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
  2043
      by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2044
    show "x\<in>cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2045
      unfolding mem_box
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2046
    proof safe
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2047
      fix i :: 'a
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2048
      assume i: "i \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2049
      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
  2050
        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
  2051
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2052
  next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2053
    fix x
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2054
    assume x: "x \<in> cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2055
    have "\<forall>i\<in>Basis.
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2056
      \<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
  2057
      (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
  2058
      unfolding mem_box
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2059
    proof
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2060
      fix i :: 'a
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2061
      assume i: "i \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2062
      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
  2063
        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
  2064
      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
  2065
        by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2066
    qed
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  2067
    then obtain \<alpha> \<beta> where
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  2068
     "\<forall>i\<in>Basis. (\<alpha> \<bullet> i = a \<bullet> i \<and> \<beta> \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<or>
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  2069
         \<alpha> \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<and> \<beta> \<bullet> i = b \<bullet> i) \<and> \<alpha> \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> \<beta> \<bullet> i"
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  2070
      by (auto simp: choice_Basis_iff)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2071
    then show "x\<in>\<Union>?A"
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  2072
      by (force simp add: mem_box)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2073
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2074
  finally show False
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2075
    using assms by auto
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
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2078
lemma interval_bisection:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2079
  fixes type :: "'a::euclidean_space"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2080
  assumes "P {}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2081
    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
  2082
    and "\<not> P (cbox a (b::'a))"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2083
  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
  2084
    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
  2085
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2086
  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
  2087
    (\<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
  2088
       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
  2089
  proof
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2090
    show "?P x" for x
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2091
    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
  2092
      case True
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2093
      then show ?thesis by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2094
    next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2095
      case as: False
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2096
      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
  2097
        "\<forall>i\<in>Basis.
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2098
           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
  2099
           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
  2100
           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
  2101
           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
  2102
        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
  2103
      then show ?thesis
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  2104
        by (rule_tac x="(c,d)" in exI) auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2105
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2106
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2107
  then obtain f where f:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2108
    "\<forall>x.
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2109
      \<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
  2110
      \<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
  2111
        (\<forall>i\<in>Basis.
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2112
            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
  2113
            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
  2114
            snd (f x) \<bullet> i \<le> snd x \<bullet> i \<and>
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  2115
            2 * (snd (f x) \<bullet> i - fst (f x) \<bullet> i) \<le> snd x \<bullet> i - fst x \<bullet> i)" by metis
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2116
  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
  2117
  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
  2118
    (\<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
  2119
    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
  2120
  proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2121
    show "A 0 = a" "B 0 = b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2122
      unfolding ab_def by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2123
    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
  2124
    show "?P n" for n
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2125
    proof (induct n)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2126
      case 0
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2127
      then show ?case
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  2128
        unfolding S using \<open>\<not> P (cbox a b)\<close> f by auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2129
    next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2130
      case (Suc n)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2131
      show ?case
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2132
        unfolding S
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2133
        apply (rule f[rule_format])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2134
        using Suc
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2135
        unfolding S
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2136
        apply auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2137
        done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2138
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2139
  qed
66193
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2140
  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
  2141
                 "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
  2142
                if "i\<in>Basis" for i n
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2143
    using that by blast+
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2144
  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
  2145
    using ABRAW by blast
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2146
  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
  2147
    if e: "0 < e" for e
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2148
  proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2149
    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
  2150
      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
  2151
    show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2152
    proof (rule exI [where x=n], clarify)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2153
      fix x y
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2154
      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
  2155
      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
  2156
        unfolding dist_norm by(rule norm_le_l1)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
  2157
      also have "\<dots> \<le> sum (\<lambda>i. B n\<bullet>i - A n\<bullet>i) Basis"
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
  2158
      proof (rule sum_mono)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2159
        fix i :: 'a
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2160
        assume i: "i \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2161
        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
  2162
          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
  2163
          by (auto simp: inner_diff_left)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2164
      qed
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
  2165
      also have "\<dots> \<le> sum (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis / 2^n"
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
  2166
        unfolding sum_divide_distrib
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
  2167
      proof (rule sum_mono)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2168
        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
  2169
        proof (induct n)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2170
          case 0
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2171
          then show ?case
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2172
            unfolding AB by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2173
        next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2174
          case (Suc n)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2175
          have "B (Suc n) \<bullet> i - A (Suc n) \<bullet> i \<le> (B n \<bullet> i - A n \<bullet> i) / 2"
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  2176
            using AB(3) that AB(4)[of i n] using i by auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2177
          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
  2178
            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
  2179
          finally show ?case .
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2180
        qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2181
      qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2182
      also have "\<dots> < e"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2183
        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
  2184
      finally show "dist x y < e" .
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2185
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2186
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2187
  {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2188
    fix n m :: nat
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2189
    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
  2190
    proof (induction rule: inc_induct)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2191
      case (step i)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2192
      show ?case
66193
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2193
        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
  2194
    qed simp
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2195
  } note ABsubset = this
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2196
  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
  2197
  proof (rule decreasing_closed_nest)
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2198
    show "\<forall>n. closed (cbox (A n) (B n))"
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2199
      by (simp add: closed_cbox)
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2200
    show "\<forall>n. cbox (A n) (B n) \<noteq> {}"
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2201
      by (meson AB dual_order.trans interval_not_empty)
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2202
  qed (use ABsubset interv in auto)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2203
  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
  2204
    by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2205
  show thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2206
  proof (rule that[rule_format, of x0])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2207
    show "x0\<in>cbox a b"
66193
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2208
      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
  2209
    fix e :: real
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2210
    assume "e > 0"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2211
    from interv[OF this] obtain n
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2212
      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
  2213
    have "\<not> P (cbox (A n) (B n))"
66193
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2214
    proof (cases "0 < n")
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2215
      case True then show ?thesis 
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2216
        by (metis Suc_pred' notPAB) 
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2217
    next
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2218
      case False then show ?thesis
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2219
        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
  2220
    qed
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2221
    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
  2222
      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
  2223
    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
  2224
      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
  2225
    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
  2226
      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
  2227
      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
  2228
      apply (auto simp: x0)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2229
      done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2230
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2231
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2232
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2233
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2234
subsection \<open>Cousin's lemma.\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2235
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2236
lemma fine_division_exists:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2237
  fixes a b :: "'a::euclidean_space"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2238
  assumes "gauge g"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2239
  obtains p where "p tagged_division_of (cbox a b)" "g fine p"
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  2240
proof (cases "\<exists>p. p tagged_division_of (cbox a b) \<and> g fine p")
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  2241
  case True
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  2242
  then show ?thesis
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  2243
    using that by auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2244
next
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  2245
  case False
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  2246
  assume "\<not> (\<exists>p. p tagged_division_of (cbox a b) \<and> g fine p)"
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2247
  obtain x where x:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2248
      "x \<in> (cbox a b)"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2249
      "\<And>e. 0 < e \<Longrightarrow>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2250
        \<exists>c d.
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2251
          x \<in> cbox c d \<and>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2252
          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
  2253
          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
  2254
          \<not> (\<exists>p. p tagged_division_of cbox c d \<and> g fine p)"
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  2255
    apply (rule interval_bisection[of "\<lambda>s. \<exists>p. p tagged_division_of s \<and> g fine p", OF _ _ False])
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2256
    apply (simp add: fine_def)
66497
18a6478a574c More tidying, and renaming of theorems
paulson <lp15@cam.ac.uk>
parents: 66365
diff changeset
  2257
    apply (metis tagged_division_Un fine_Un)
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  2258
    apply auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2259
    done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2260
  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
  2261
    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
  2262
  from x(2)[OF e(1)]
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2263
  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
  2264
                        "cbox c d \<subseteq> ball x e"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2265
                        "cbox c d \<subseteq> cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2266
                        "\<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
  2267
    by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2268
  have "g fine {(x, cbox c d)}"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2269
    unfolding fine_def using e using c_d(2) by auto
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  2270
  then show ?thesis
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2271
    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
  2272
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2273
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2274
lemma fine_division_exists_real:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2275
  fixes a b :: real
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2276
  assumes "gauge g"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2277
  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
  2278
  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
  2279
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2280
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
  2281
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2282
lemma tagged_division_finer:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2283
  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
  2284
  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
  2285
    and "gauge d"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2286
  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
  2287
    and "d fine q"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2288
    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
  2289
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2290
  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
  2291
    (\<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
  2292
      (\<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
  2293
  {
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2294
    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
  2295
      using assms(1)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2296
      unfolding tagged_division_of_def
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2297
      by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2298
    presume "\<And>p. finite p \<Longrightarrow> ?P p"
66300
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  2299
    from this[rule_format,OF * assms(2)] 
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  2300
    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
  2301
      by auto
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  2302
    with that[of q] show ?thesis
829f1f62b087 more cleanup of Tagged_Division
paulson <lp15@cam.ac.uk>
parents: 66299
diff changeset
  2303
      using assms(1) by auto
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2304
  }
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2305
  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
  2306
  assume as: "finite p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2307
  show "?P p"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2308
    apply rule
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2309
    apply rule
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2310
    using as
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2311
  proof (induct p)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2312
    case empty
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2313
    show ?case
66318
f2e1047d6cc1 more tidying
paulson <lp15@cam.ac.uk>
parents: 66317
diff changeset
  2314
      by (force simp add: fine_def)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2315
  next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2316
    case (insert xk p)
66314
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2317
    obtain x k where xk: "xk = (x, k)"
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2318
      using surj_pair[of xk] by metis 
66193
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2319
    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
  2320
                and "d fine q1"
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2321
                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
  2322
      using case_prodD tagged_partial_division_subset[OF insert(4) subset_insertI]
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2323
      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
  2324
    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
  2325
      unfolding xk by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2326
    note p = tagged_partial_division_ofD[OF insert(4)]
66314
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2327
    obtain u v where uv: "k = cbox u v"
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2328
      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
  2329
    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
  2330
      apply (rule finite_subset[of _ "snd ` p"])
66314
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2331
      using image_iff apply fastforce
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2332
      using insert.hyps(1) by blast
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2333
    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
  2334
    proof (rule Int_interior_Union_intervals)
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2335
      show "open (interior (cbox u v))"
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2336
        by auto
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2337
      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
  2338
        using p(4) by auto
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2339
      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
  2340
        by clarify (metis insert.hyps(2) insert_iff interior_cbox p(5) uv xk)
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2341
    qed
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2342
    show ?case
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2343
    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
  2344
      case True
66314
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2345
      have "{(x, cbox u v)} tagged_division_of cbox u v"
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2346
        by (simp add: p(2) uv xk tagged_division_of_self)
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2347
      then have "{(x, cbox u v)} \<union> q1 tagged_division_of \<Union>{k. \<exists>x. (x, k) \<in> insert xk p}"
66497
18a6478a574c More tidying, and renaming of theorems
paulson <lp15@cam.ac.uk>
parents: 66365
diff changeset
  2348
        unfolding * uv by (metis (no_types, lifting) int q1 tagged_division_Un)
66314
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2349
      with True show ?thesis
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2350
        apply (rule_tac x="{(x,cbox u v)} \<union> q1" in exI)
66314
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2351
        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
  2352
        done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2353
    next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2354
      case False
66314
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2355
      obtain q2 where q2: "q2 tagged_division_of cbox u v" "d fine q2"
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2356
        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
  2357
      show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2358
        apply (rule_tac x="q2 \<union> q1" in exI)
66314
52914a618299 removed all "guess"
paulson
parents: 66306
diff changeset
  2359
        apply (intro conjI)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2360
        unfolding * uv
66497
18a6478a574c More tidying, and renaming of theorems
paulson <lp15@cam.ac.uk>
parents: 66365
diff changeset
  2361
        apply (rule tagged_division_Un q2 q1 int fine_Un)+
66193
6e6eeef63589 More tidying of horrible proofs
paulson <lp15@cam.ac.uk>
parents: 66192
diff changeset
  2362
          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
  2363
        done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2364
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2365
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2366
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2367
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2368
subsubsection \<open>Covering lemma\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2369
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2370
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
  2371
  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
  2372
  "Introduction to Gauge Integrals". \<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2373
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2374
proposition covering_lemma:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2375
  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
  2376
  obtains \<D> where
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2377
    "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
  2378
    "\<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
  2379
    "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
  2380
    "\<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
  2381
    "\<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
  2382
    "S \<subseteq> \<Union>\<D>"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2383
proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2384
  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
  2385
    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
  2386
  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
  2387
                    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
  2388
                         (\<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
  2389
  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
  2390
  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
  2391
             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
  2392
             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
  2393
             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
  2394
             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
  2395
             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
  2396
             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
  2397
  proof
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2398
    show "countable ?D0"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2399
      by (simp add: countable_PiE)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2400
  next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2401
    show "\<Union>?D0 \<subseteq> cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2402
      apply (simp add: UN_subset_iff)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2403
      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
  2404
       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
  2405
      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
  2406
      done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2407
  next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2408
    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
  2409
      using \<open>box a b \<noteq> {}\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2410
      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
  2411
  next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2412
    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
  2413
      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
  2414
    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
  2415
      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
  2416
    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
  2417
      fix v w m and n::nat
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64910
diff changeset
  2418
      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
  2419
      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
  2420
        apply (simp add: subset_box disjoint_interval)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2421
        apply (rule ccontr)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2422
        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
  2423
        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
  2424
        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
  2425
        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
  2426
        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
  2427
        done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2428
      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
  2429
        by meson
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2430
    qed auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2431
    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
  2432
      apply (erule imageE SigmaE)+
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2433
      using * by simp
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2434
  next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2435
    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
  2436
    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
  2437
      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
  2438
      proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2439
        obtain e where "0 < e"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2440
                   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
  2441
        proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2442
          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
  2443
            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
  2444
          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
  2445
            using openE by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2446
          have "norm (x - y) < \<epsilon>"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2447
               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
  2448
          proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2449
            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
  2450
              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
  2451
            also have "... \<le> DIM('a) * (\<epsilon> / (2 * real DIM('a)))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
  2452
              by (meson sum_bounded_above that)
63957
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2453
            also have "... = \<epsilon> / 2"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2454
              by (simp add: divide_simps)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2455
            also have "... < \<epsilon>"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2456
              by (simp add: \<open>0 < \<epsilon>\<close>)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2457
            finally show ?thesis .
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2458
          qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2459
          then show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2460
            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
  2461
        qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2462
        have xab: "x \<in> cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2463
          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
  2464
        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
  2465
          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
  2466
          by (auto simp: divide_simps)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2467
        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
  2468
          by (auto simp: divide_simps)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2469
        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
  2470
        proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2471
          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
  2472
            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
  2473
          also have "... < e * 2 ^ n"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2474
            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
  2475
          finally show ?thesis .
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2476
        qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2477
        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
  2478
                 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
  2479
          by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2480
        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
  2481
               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
  2482
        proof
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2483
          fix i::'a assume "i \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2484
          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
  2485
            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
  2486
          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
  2487
                          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
  2488
          proof cases
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2489
            case 1 then show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2490
              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
  2491
          next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2492
            case 2
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2493
            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
  2494
              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
  2495
            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
  2496
            show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2497
            proof (intro exI conjI)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2498
              show "?k < 2 ^ n"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2499
                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
  2500
                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
  2501
            next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2502
              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
  2503
                    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
  2504
                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
  2505
                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
  2506
                  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
  2507
                done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2508
              also have "... = x \<bullet> i"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2509
                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
  2510
              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
  2511
            next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2512
              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
  2513
                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
  2514
              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
  2515
                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
  2516
                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
  2517
                  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
  2518
                done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2519
              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
  2520
            qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2521
          qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2522
        qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2523
        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
  2524
          apply (simp add: mem_box Bex_def)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2525
          apply (clarify dest!: bchoice)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2526
          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
  2527
          done
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2528
        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
  2529
          apply (clarsimp simp add: mem_box)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2530
          apply (rule e)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2531
          apply (drule bspec D, assumption)+
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2532
          apply (erule order_trans)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2533
          apply (simp add: divide_simps)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2534
          using bai by (force simp: algebra_simps)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2535
        ultimately show ?thesis by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2536
      qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2537
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2538
  next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2539
    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
  2540
      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
  2541
  next
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2542
    obtain j::'a where "j \<in> Basis"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2543
      using nonempty_Basis by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2544
    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
  2545
    proof (rule finite_subset)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2546
      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
  2547
                                        (\<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
  2548
                ` (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
  2549
      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
  2550
      proof -
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2551
        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
  2552
                  \<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
  2553
          by (auto simp: divide_simps algebra_simps)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2554
        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
  2555
          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
  2556
        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
  2557
              (((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
  2558
          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
  2559
        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
  2560
          ((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
  2561
          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
  2562
        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
  2563
          by (rule dd)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2564
        then have "m \<le> n"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2565
          by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2566
        show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2567
          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
  2568
      qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2569
      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
  2570
        by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2571
      show "finite ?B"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2572
        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
  2573
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2574
    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
  2575
      using that by auto
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2576
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2577
  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
  2578
  obtain \<D> where count: "countable \<D>"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2579
             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
  2580
             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
  2581
             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
  2582
             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
  2583
             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
  2584
             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
  2585
  proof
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2586
    show "countable ?D1" using count countable_subset
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2587
      by (simp add: count countable_subset)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2588
    show "\<Union>?D1 \<subseteq> cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2589
      using sub by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2590
    show "S \<subseteq> \<Union>?D1"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2591
      using SK by (force simp:)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2592
    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
  2593
      using int by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2594
    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
  2595
      using intdj by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2596
    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
  2597
      by auto
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> ?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
  2599
      using cbox by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2600
    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
  2601
      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
  2602
  qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2603
  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
  2604
  show ?thesis
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2605
  proof (rule that)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2606
    show "countable ?\<D>"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2607
      by (blast intro: countable_subset [OF _ count])
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2608
    show "\<Union>?\<D> \<subseteq> cbox a b"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2609
      using sub by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2610
    show "S \<subseteq> \<Union>?\<D>"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2611
    proof clarsimp
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2612
      fix x
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2613
      assume "x \<in> S"
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2614
      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
  2615
      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
  2616
      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
  2617
      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
  2618
      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
  2619
        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
  2620
      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
  2621
        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
  2622
      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
  2623
        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
  2624
      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
  2625
        by blast
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2626
    qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2627
    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
  2628
      by (blast intro: dest: int)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2629
    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
  2630
      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
  2631
    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
  2632
      using SK by force
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2633
    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
  2634
      using cbox by force
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
qed
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2637
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2638
subsection \<open>Division filter\<close>
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2639
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2640
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
  2641
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2642
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
  2643
  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
  2644
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2645
lemma eventually_division_filter:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2646
  "(\<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
  2647
    (\<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
  2648
  unfolding division_filter_def
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2649
proof (subst eventually_INF_base; clarsimp)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2650
  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
  2651
    {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
  2652
    {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
  2653
    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
  2654
qed (auto simp: eventually_principal)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2655
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2656
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
  2657
  unfolding trivial_limit_def eventually_division_filter
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2658
  by (auto elim: fine_division_exists)
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2659
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2660
lemma eventually_division_filter_tagged_division:
c3da799b1b45 HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
diff changeset
  2661
  "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
  2662
  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
  2663
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63957
diff changeset
  2664
end