src/HOL/Analysis/Ordered_Euclidean_Space.thy
author paulson <lp15@cam.ac.uk>
Fri, 26 Apr 2019 16:51:40 +0100
changeset 70196 b7ef9090feed
parent 70136 f03a01a18c6e
child 70378 ebd108578ab1
permissions -rw-r--r--
Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69874
diff changeset
     1
section \<open>Ordered Euclidean Space\<close>
69631
6c3e6038e74c tuned headers
nipkow
parents: 69508
diff changeset
     2
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
     3
theory Ordered_Euclidean_Space
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
     4
imports
56189
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
     5
  Cartesian_Euclidean_Space
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 64758
diff changeset
     6
  "HOL-Library.Product_Order"
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
     7
begin
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
     8
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
     9
text \<open>An ordering on euclidean spaces that will allow us to talk about intervals\<close>
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    10
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    11
class ordered_euclidean_space = ord + inf + sup + abs + Inf + Sup + euclidean_space +
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    12
  assumes eucl_le: "x \<le> y \<longleftrightarrow> (\<forall>i\<in>Basis. x \<bullet> i \<le> y \<bullet> i)"
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    13
  assumes eucl_less_le_not_le: "x < y \<longleftrightarrow> x \<le> y \<and> \<not> y \<le> x"
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    14
  assumes eucl_inf: "inf x y = (\<Sum>i\<in>Basis. inf (x \<bullet> i) (y \<bullet> i) *\<^sub>R i)"
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    15
  assumes eucl_sup: "sup x y = (\<Sum>i\<in>Basis. sup (x \<bullet> i) (y \<bullet> i) *\<^sub>R i)"
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69000
diff changeset
    16
  assumes eucl_Inf: "Inf X = (\<Sum>i\<in>Basis. (INF x\<in>X. x \<bullet> i) *\<^sub>R i)"
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69000
diff changeset
    17
  assumes eucl_Sup: "Sup X = (\<Sum>i\<in>Basis. (SUP x\<in>X. x \<bullet> i) *\<^sub>R i)"
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61808
diff changeset
    18
  assumes eucl_abs: "\<bar>x\<bar> = (\<Sum>i\<in>Basis. \<bar>x \<bullet> i\<bar> *\<^sub>R i)"
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    19
begin
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    20
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    21
subclass order
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
    22
  by standard
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    23
    (auto simp: eucl_le eucl_less_le_not_le intro!: euclidean_eqI antisym intro: order.trans)
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    24
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    25
subclass ordered_ab_group_add_abs
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
    26
  by standard (auto simp: eucl_le inner_add_left eucl_abs abs_leI)
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    27
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    28
subclass ordered_real_vector
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
    29
  by standard (auto simp: eucl_le intro!: mult_left_mono mult_right_mono)
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    30
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    31
subclass lattice
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
    32
  by standard (auto simp: eucl_inf eucl_sup eucl_le)
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    33
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    34
subclass distrib_lattice
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
    35
  by standard (auto simp: eucl_inf eucl_sup sup_inf_distrib1 intro!: euclidean_eqI)
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    36
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    37
subclass conditionally_complete_lattice
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
    38
proof
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    39
  fix z::'a and X::"'a set"
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    40
  assume "X \<noteq> {}"
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    41
  hence "\<And>i. (\<lambda>x. x \<bullet> i) ` X \<noteq> {}" by simp
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    42
  thus "(\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> z \<le> Inf X" "(\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X \<le> z"
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 61945
diff changeset
    43
    by (auto simp: eucl_Inf eucl_Sup eucl_le
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    44
      intro!: cInf_greatest cSup_least)
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    45
qed (force intro!: cInf_lower cSup_upper
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    46
      simp: bdd_below_def bdd_above_def preorder_class.bdd_below_def preorder_class.bdd_above_def
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 61945
diff changeset
    47
        eucl_Inf eucl_Sup eucl_le)+
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    48
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
    49
lemma inner_Basis_inf_left: "i \<in> Basis \<Longrightarrow> inf x y \<bullet> i = inf (x \<bullet> i) (y \<bullet> i)"
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    50
  and inner_Basis_sup_left: "i \<in> Basis \<Longrightarrow> sup x y \<bullet> i = sup (x \<bullet> i) (y \<bullet> i)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63886
diff changeset
    51
  by (simp_all add: eucl_inf eucl_sup inner_sum_left inner_Basis if_distrib comm_monoid_add_class.sum.delta
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    52
      cong: if_cong)
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    53
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
    54
lemma inner_Basis_INF_left: "i \<in> Basis \<Longrightarrow> (INF x\<in>X. f x) \<bullet> i = (INF x\<in>X. f x \<bullet> i)"
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69000
diff changeset
    55
  and inner_Basis_SUP_left: "i \<in> Basis \<Longrightarrow> (SUP x\<in>X. f x) \<bullet> i = (SUP x\<in>X. f x \<bullet> i)"
69861
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69730
diff changeset
    56
  using eucl_Sup [of "f ` X"] eucl_Inf [of "f ` X"] by (simp_all add: image_comp)
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    57
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
    58
lemma abs_inner: "i \<in> Basis \<Longrightarrow> \<bar>x\<bar> \<bullet> i = \<bar>x \<bullet> i\<bar>"
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    59
  by (auto simp: eucl_abs)
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    60
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
    61
lemma
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    62
  abs_scaleR: "\<bar>a *\<^sub>R b\<bar> = \<bar>a\<bar> *\<^sub>R \<bar>b\<bar>"
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    63
  by (auto simp: eucl_abs abs_mult intro!: euclidean_eqI)
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    64
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
    65
lemma interval_inner_leI:
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    66
  assumes "x \<in> {a .. b}" "0 \<le> i"
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    67
  shows "a\<bullet>i \<le> x\<bullet>i" "x\<bullet>i \<le> b\<bullet>i"
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    68
  using assms
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    69
  unfolding euclidean_inner[of a i] euclidean_inner[of x i] euclidean_inner[of b i]
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63886
diff changeset
    70
  by (auto intro!: ordered_comm_monoid_add_class.sum_mono mult_right_mono simp: eucl_le)
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    71
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
    72
lemma inner_nonneg_nonneg:
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    73
  shows "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> a \<bullet> b"
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    74
  using interval_inner_leI[of a 0 a b]
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    75
  by auto
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    76
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
    77
lemma inner_Basis_mono:
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    78
  shows "a \<le> b \<Longrightarrow> c \<in> Basis  \<Longrightarrow> a \<bullet> c \<le> b \<bullet> c"
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    79
  by (simp add: eucl_le)
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    80
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
    81
lemma Basis_nonneg[intro, simp]: "i \<in> Basis \<Longrightarrow> 0 \<le> i"
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    82
  by (auto simp: eucl_le inner_Basis)
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
    83
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
    84
lemma Sup_eq_maximum_componentwise:
54781
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
    85
  fixes s::"'a set"
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
    86
  assumes i: "\<And>b. b \<in> Basis \<Longrightarrow> X \<bullet> b = i b \<bullet> b"
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
    87
  assumes sup: "\<And>b x. b \<in> Basis \<Longrightarrow> x \<in> s \<Longrightarrow> x \<bullet> b \<le> X \<bullet> b"
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
    88
  assumes i_s: "\<And>b. b \<in> Basis \<Longrightarrow> (i b \<bullet> b) \<in> (\<lambda>x. x \<bullet> b) ` s"
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
    89
  shows "Sup s = X"
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
    90
  using assms
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63886
diff changeset
    91
  unfolding eucl_Sup euclidean_representation_sum
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 61945
diff changeset
    92
  by (auto intro!: conditionally_complete_lattice_class.cSup_eq_maximum)
54781
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
    93
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
    94
lemma Inf_eq_minimum_componentwise:
54781
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
    95
  assumes i: "\<And>b. b \<in> Basis \<Longrightarrow> X \<bullet> b = i b \<bullet> b"
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
    96
  assumes sup: "\<And>b x. b \<in> Basis \<Longrightarrow> x \<in> s \<Longrightarrow> X \<bullet> b \<le> x \<bullet> b"
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
    97
  assumes i_s: "\<And>b. b \<in> Basis \<Longrightarrow> (i b \<bullet> b) \<in> (\<lambda>x. x \<bullet> b) ` s"
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
    98
  shows "Inf s = X"
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
    99
  using assms
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63886
diff changeset
   100
  unfolding eucl_Inf euclidean_representation_sum
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 61945
diff changeset
   101
  by (auto intro!: conditionally_complete_lattice_class.cInf_eq_minimum)
54781
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   102
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
   103
end
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
   104
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   105
proposition  compact_attains_Inf_componentwise:
54781
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   106
  fixes b::"'a::ordered_euclidean_space"
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   107
  assumes "b \<in> Basis" assumes "X \<noteq> {}" "compact X"
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   108
  obtains x where "x \<in> X" "x \<bullet> b = Inf X \<bullet> b" "\<And>y. y \<in> X \<Longrightarrow> x \<bullet> b \<le> y \<bullet> b"
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   109
proof atomize_elim
54781
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   110
  let ?proj = "(\<lambda>x. x \<bullet> b) ` X"
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   111
  from assms have "compact ?proj" "?proj \<noteq> {}"
56371
fb9ae0727548 extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents: 56190
diff changeset
   112
    by (auto intro!: compact_continuous_image continuous_intros)
54781
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   113
  from compact_attains_inf[OF this]
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   114
  obtain s x
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   115
    where s: "s\<in>(\<lambda>x. x \<bullet> b) ` X" "\<And>t. t\<in>(\<lambda>x. x \<bullet> b) ` X \<Longrightarrow> s \<le> t"
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   116
      and x: "x \<in> X" "s = x \<bullet> b" "\<And>y. y \<in> X \<Longrightarrow> x \<bullet> b \<le> y \<bullet> b"
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   117
    by auto
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   118
  hence "Inf ?proj = x \<bullet> b"
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 61945
diff changeset
   119
    by (auto intro!: conditionally_complete_lattice_class.cInf_eq_minimum)
54781
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   120
  hence "x \<bullet> b = Inf X \<bullet> b"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63886
diff changeset
   121
    by (auto simp: eucl_Inf inner_sum_left inner_Basis if_distrib \<open>b \<in> Basis\<close> sum.delta
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 55413
diff changeset
   122
      cong: if_cong)
54781
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   123
  with x show "\<exists>x. x \<in> X \<and> x \<bullet> b = Inf X \<bullet> b \<and> (\<forall>y. y \<in> X \<longrightarrow> x \<bullet> b \<le> y \<bullet> b)" by blast
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   124
qed
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   125
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   126
proposition
54781
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   127
  compact_attains_Sup_componentwise:
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   128
  fixes b::"'a::ordered_euclidean_space"
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   129
  assumes "b \<in> Basis" assumes "X \<noteq> {}" "compact X"
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   130
  obtains x where "x \<in> X" "x \<bullet> b = Sup X \<bullet> b" "\<And>y. y \<in> X \<Longrightarrow> y \<bullet> b \<le> x \<bullet> b"
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   131
proof atomize_elim
54781
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   132
  let ?proj = "(\<lambda>x. x \<bullet> b) ` X"
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   133
  from assms have "compact ?proj" "?proj \<noteq> {}"
56371
fb9ae0727548 extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents: 56190
diff changeset
   134
    by (auto intro!: compact_continuous_image continuous_intros)
54781
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   135
  from compact_attains_sup[OF this]
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   136
  obtain s x
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   137
    where s: "s\<in>(\<lambda>x. x \<bullet> b) ` X" "\<And>t. t\<in>(\<lambda>x. x \<bullet> b) ` X \<Longrightarrow> t \<le> s"
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   138
      and x: "x \<in> X" "s = x \<bullet> b" "\<And>y. y \<in> X \<Longrightarrow> y \<bullet> b \<le> x \<bullet> b"
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   139
    by auto
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   140
  hence "Sup ?proj = x \<bullet> b"
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 61945
diff changeset
   141
    by (auto intro!: cSup_eq_maximum)
54781
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   142
  hence "x \<bullet> b = Sup X \<bullet> b"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63886
diff changeset
   143
    by (auto simp: eucl_Sup[where 'a='a] inner_sum_left inner_Basis if_distrib \<open>b \<in> Basis\<close> sum.delta
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 61945
diff changeset
   144
      cong: if_cong)
54781
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   145
  with x show "\<exists>x. x \<in> X \<and> x \<bullet> b = Sup X \<bullet> b \<and> (\<forall>y. y \<in> X \<longrightarrow> y \<bullet> b \<le> x \<bullet> b)" by blast
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   146
qed
fe462aa28c3d additional lemmas
immler
parents: 54780
diff changeset
   147
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   148
lemma (in order) atLeastatMost_empty'[simp]:
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69260
diff changeset
   149
  "(\<not> a \<le> b) \<Longrightarrow> {a..b} = {}"
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
   150
  by (auto)
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
   151
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
   152
instance real :: ordered_euclidean_space
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 61945
diff changeset
   153
  by standard auto
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
   154
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   155
lemma in_Basis_prod_iff:
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
   156
  fixes i::"'a::euclidean_space*'b::euclidean_space"
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
   157
  shows "i \<in> Basis \<longleftrightarrow> fst i = 0 \<and> snd i \<in> Basis \<or> snd i = 0 \<and> fst i \<in> Basis"
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
   158
  by (cases i) (auto simp: Basis_prod_def)
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
   159
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
   160
instantiation\<^marker>\<open>tag unimportant\<close> prod :: (abs, abs) abs
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
   161
begin
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
   162
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61808
diff changeset
   163
definition "\<bar>x\<bar> = (\<bar>fst x\<bar>, \<bar>snd x\<bar>)"
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
   164
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61808
diff changeset
   165
instance ..
1135b8de26c3 more symbols;
wenzelm
parents: 61808
diff changeset
   166
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
   167
end
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
   168
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
   169
instance prod :: (ordered_euclidean_space, ordered_euclidean_space) ordered_euclidean_space
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
   170
  by standard
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63886
diff changeset
   171
    (auto intro!: add_mono simp add: euclidean_representation_sum'  Ball_def inner_prod_def
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
   172
      in_Basis_prod_iff inner_Basis_inf_left inner_Basis_sup_left inner_Basis_INF_left Inf_prod_def
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
   173
      inner_Basis_SUP_left Sup_prod_def less_prod_def less_eq_prod_def eucl_le[where 'a='a]
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
   174
      eucl_le[where 'a='b] abs_prod_def abs_inner)
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
   175
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   176
text\<open>Instantiation for intervals on \<open>ordered_euclidean_space\<close>\<close>
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56166
diff changeset
   177
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   178
proposition
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60420
diff changeset
   179
  fixes a :: "'a::ordered_euclidean_space"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56166
diff changeset
   180
  shows cbox_interval: "cbox a b = {a..b}"
0268784f60da use cbox to relax class constraints
immler
parents: 56166
diff changeset
   181
    and interval_cbox: "{a..b} = cbox a b"
0268784f60da use cbox to relax class constraints
immler
parents: 56166
diff changeset
   182
    and eucl_le_atMost: "{x. \<forall>i\<in>Basis. x \<bullet> i <= a \<bullet> i} = {..a}"
0268784f60da use cbox to relax class constraints
immler
parents: 56166
diff changeset
   183
    and eucl_le_atLeast: "{x. \<forall>i\<in>Basis. a \<bullet> i <= x \<bullet> i} = {a..}"
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   184
  by (auto simp: eucl_le[where 'a='a] eucl_less_def box_def cbox_def)
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56166
diff changeset
   185
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   186
lemma vec_nth_real_1_iff_cbox [simp]:
67981
349c639e593c more new theorems on real^1, matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67685
diff changeset
   187
  fixes a b :: real
349c639e593c more new theorems on real^1, matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67685
diff changeset
   188
  shows "(\<lambda>x::real^1. x $ 1) ` S = {a..b} \<longleftrightarrow> S = cbox (vec a) (vec b)"
349c639e593c more new theorems on real^1, matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67685
diff changeset
   189
  by (metis interval_cbox vec_nth_1_iff_cbox)
349c639e593c more new theorems on real^1, matrices, etc.
paulson <lp15@cam.ac.uk>
parents: 67685
diff changeset
   190
70196
b7ef9090feed Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   191
lemma sums_vec_nth :
b7ef9090feed Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   192
  assumes "f sums a"
b7ef9090feed Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   193
  shows "(\<lambda>x. f x $ i) sums a $ i"
b7ef9090feed Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   194
  using assms unfolding sums_def
b7ef9090feed Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   195
  by (auto dest: tendsto_vec_nth [where i=i])
b7ef9090feed Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   196
b7ef9090feed Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   197
lemma summable_vec_nth :
b7ef9090feed Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   198
  assumes "summable f"
b7ef9090feed Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   199
  shows "summable (\<lambda>x. f x $ i)"
b7ef9090feed Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   200
  using assms unfolding summable_def
b7ef9090feed Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   201
  by (blast intro: sums_vec_nth)
b7ef9090feed Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   202
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   203
lemma closed_eucl_atLeastAtMost[simp, intro]:
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60420
diff changeset
   204
  fixes a :: "'a::ordered_euclidean_space"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56166
diff changeset
   205
  shows "closed {a..b}"
0268784f60da use cbox to relax class constraints
immler
parents: 56166
diff changeset
   206
  by (simp add: cbox_interval[symmetric] closed_cbox)
0268784f60da use cbox to relax class constraints
immler
parents: 56166
diff changeset
   207
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   208
lemma closed_eucl_atMost[simp, intro]:
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60420
diff changeset
   209
  fixes a :: "'a::ordered_euclidean_space"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56166
diff changeset
   210
  shows "closed {..a}"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   211
  by (simp add: closed_interval_left eucl_le_atMost[symmetric])
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56166
diff changeset
   212
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   213
lemma closed_eucl_atLeast[simp, intro]:
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60420
diff changeset
   214
  fixes a :: "'a::ordered_euclidean_space"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56166
diff changeset
   215
  shows "closed {a..}"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   216
  by (simp add: closed_interval_right eucl_le_atLeast[symmetric])
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56166
diff changeset
   217
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   218
lemma bounded_closed_interval [simp]:
56189
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
   219
  fixes a :: "'a::ordered_euclidean_space"
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
   220
  shows "bounded {a .. b}"
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
   221
  using bounded_cbox[of a b]
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
   222
  by (metis interval_cbox)
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
   223
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   224
lemma convex_closed_interval [simp]:
56190
f0d2609c4cdc additional lemmas
immler
parents: 56189
diff changeset
   225
  fixes a :: "'a::ordered_euclidean_space"
f0d2609c4cdc additional lemmas
immler
parents: 56189
diff changeset
   226
  shows "convex {a .. b}"
f0d2609c4cdc additional lemmas
immler
parents: 56189
diff changeset
   227
  using convex_box[of a b]
f0d2609c4cdc additional lemmas
immler
parents: 56189
diff changeset
   228
  by (metis interval_cbox)
f0d2609c4cdc additional lemmas
immler
parents: 56189
diff changeset
   229
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   230
lemma image_smult_interval:"(\<lambda>x. m *\<^sub>R (x::_::ordered_euclidean_space)) ` {a .. b} =
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56166
diff changeset
   231
  (if {a .. b} = {} then {} else if 0 \<le> m then {m *\<^sub>R a .. m *\<^sub>R b} else {m *\<^sub>R b .. m *\<^sub>R a})"
0268784f60da use cbox to relax class constraints
immler
parents: 56166
diff changeset
   232
  using image_smult_cbox[of m a b]
0268784f60da use cbox to relax class constraints
immler
parents: 56166
diff changeset
   233
  by (simp add: cbox_interval)
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
   234
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   235
lemma [simp]:
70097
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
   236
  fixes a b::"'a::ordered_euclidean_space"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
   237
  shows is_interval_ic: "is_interval {..a}"
67685
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 66827
diff changeset
   238
    and is_interval_ci: "is_interval {a..}"
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 66827
diff changeset
   239
    and is_interval_cc: "is_interval {b..a}"
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 66827
diff changeset
   240
  by (force simp: is_interval_def eucl_le[where 'a='a])+
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 66827
diff changeset
   241
69000
7cb3ddd60fd6 more lemmas
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   242
lemma connected_interval [simp]:
7cb3ddd60fd6 more lemmas
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   243
  fixes a b::"'a::ordered_euclidean_space"
7cb3ddd60fd6 more lemmas
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   244
  shows "connected {a..b}"
7cb3ddd60fd6 more lemmas
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   245
  using is_interval_cc is_interval_connected by blast
7cb3ddd60fd6 more lemmas
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   246
7cb3ddd60fd6 more lemmas
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   247
lemma path_connected_interval [simp]:
7cb3ddd60fd6 more lemmas
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   248
  fixes a b::"'a::ordered_euclidean_space"
7cb3ddd60fd6 more lemmas
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   249
  shows "path_connected {a..b}"
7cb3ddd60fd6 more lemmas
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   250
  using is_interval_cc is_interval_path_connected by blast
7cb3ddd60fd6 more lemmas
paulson <lp15@cam.ac.uk>
parents: 68833
diff changeset
   251
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69861
diff changeset
   252
lemma path_connected_Ioi[simp]: "path_connected {a<..}" for a :: real
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69861
diff changeset
   253
  by (simp add: convex_imp_path_connected)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69861
diff changeset
   254
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69861
diff changeset
   255
lemma path_connected_Ici[simp]: "path_connected {a..}" for a :: real
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69861
diff changeset
   256
  by (simp add: convex_imp_path_connected)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69861
diff changeset
   257
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69861
diff changeset
   258
lemma path_connected_Iio[simp]: "path_connected {..<a}" for a :: real
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69861
diff changeset
   259
  by (simp add: convex_imp_path_connected)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69861
diff changeset
   260
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69861
diff changeset
   261
lemma path_connected_Iic[simp]: "path_connected {..a}" for a :: real
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69861
diff changeset
   262
  by (simp add: convex_imp_path_connected)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69861
diff changeset
   263
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69861
diff changeset
   264
lemma path_connected_Ioo[simp]: "path_connected {a<..<b}" for a b :: real
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69861
diff changeset
   265
  by (simp add: convex_imp_path_connected)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69861
diff changeset
   266
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69861
diff changeset
   267
lemma path_connected_Ioc[simp]: "path_connected {a<..b}" for a b :: real
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69861
diff changeset
   268
  by (simp add: convex_imp_path_connected)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69861
diff changeset
   269
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69861
diff changeset
   270
lemma path_connected_Ico[simp]: "path_connected {a..<b}" for a b :: real
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69861
diff changeset
   271
  by (simp add: convex_imp_path_connected)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents: 69861
diff changeset
   272
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   273
lemma is_interval_real_ereal_oo: "is_interval (real_of_ereal ` {N<..<M::ereal})"
67685
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 66827
diff changeset
   274
  by (auto simp: real_atLeastGreaterThan_eq)
56189
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
   275
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   276
lemma compact_interval [simp]:
56190
f0d2609c4cdc additional lemmas
immler
parents: 56189
diff changeset
   277
  fixes a b::"'a::ordered_euclidean_space"
f0d2609c4cdc additional lemmas
immler
parents: 56189
diff changeset
   278
  shows "compact {a .. b}"
f0d2609c4cdc additional lemmas
immler
parents: 56189
diff changeset
   279
  by (metis compact_cbox interval_cbox)
f0d2609c4cdc additional lemmas
immler
parents: 56189
diff changeset
   280
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   281
lemma homeomorphic_closed_intervals:
62620
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62376
diff changeset
   282
  fixes a :: "'a::euclidean_space" and b and c :: "'a::euclidean_space" and d
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62376
diff changeset
   283
  assumes "box a b \<noteq> {}" and "box c d \<noteq> {}"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62376
diff changeset
   284
    shows "(cbox a b) homeomorphic (cbox c d)"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62376
diff changeset
   285
apply (rule homeomorphic_convex_compact)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62376
diff changeset
   286
using assms apply auto
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62376
diff changeset
   287
done
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62376
diff changeset
   288
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   289
lemma homeomorphic_closed_intervals_real:
62620
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62376
diff changeset
   290
  fixes a::real and b and c::real and d
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62376
diff changeset
   291
  assumes "a<b" and "c<d"
68239
0764ee22a4d1 tidy up of Derivative
paulson <lp15@cam.ac.uk>
parents: 68120
diff changeset
   292
  shows "{a..b} homeomorphic {c..d}"
0764ee22a4d1 tidy up of Derivative
paulson <lp15@cam.ac.uk>
parents: 68120
diff changeset
   293
  using assms by (auto intro: homeomorphic_convex_compact)
62620
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62376
diff changeset
   294
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
   295
no_notation
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
   296
  eucl_less (infix "<e" 50)
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
   297
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   298
lemma One_nonneg: "0 \<le> (\<Sum>Basis::'a::ordered_euclidean_space)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63886
diff changeset
   299
  by (auto intro: sum_nonneg)
56189
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
   300
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   301
lemma
56190
f0d2609c4cdc additional lemmas
immler
parents: 56189
diff changeset
   302
  fixes a b::"'a::ordered_euclidean_space"
f0d2609c4cdc additional lemmas
immler
parents: 56189
diff changeset
   303
  shows bdd_above_cbox[intro, simp]: "bdd_above (cbox a b)"
f0d2609c4cdc additional lemmas
immler
parents: 56189
diff changeset
   304
    and bdd_below_cbox[intro, simp]: "bdd_below (cbox a b)"
f0d2609c4cdc additional lemmas
immler
parents: 56189
diff changeset
   305
    and bdd_above_box[intro, simp]: "bdd_above (box a b)"
f0d2609c4cdc additional lemmas
immler
parents: 56189
diff changeset
   306
    and bdd_below_box[intro, simp]: "bdd_below (box a b)"
f0d2609c4cdc additional lemmas
immler
parents: 56189
diff changeset
   307
  unfolding atomize_conj
f0d2609c4cdc additional lemmas
immler
parents: 56189
diff changeset
   308
  by (metis bdd_above_Icc bdd_above_mono bdd_below_Icc bdd_below_mono bounded_box
68120
2f161c6910f7 tidying more messy proofs
paulson <lp15@cam.ac.uk>
parents: 67981
diff changeset
   309
            bounded_subset_cbox_symmetric interval_cbox)
56190
f0d2609c4cdc additional lemmas
immler
parents: 56189
diff changeset
   310
56189
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
   311
instantiation vec :: (ordered_euclidean_space, finite) ordered_euclidean_space
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
   312
begin
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
   313
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
   314
definition\<^marker>\<open>tag important\<close> "inf x y = (\<chi> i. inf (x $ i) (y $ i))"
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
   315
definition\<^marker>\<open>tag important\<close> "sup x y = (\<chi> i. sup (x $ i) (y $ i))"
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
   316
definition\<^marker>\<open>tag important\<close> "Inf X = (\<chi> i. (INF x\<in>X. x $ i))"
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
   317
definition\<^marker>\<open>tag important\<close> "Sup X = (\<chi> i. (SUP x\<in>X. x $ i))"
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
   318
definition\<^marker>\<open>tag important\<close> "\<bar>x\<bar> = (\<chi> i. \<bar>x $ i\<bar>)"
56189
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
   319
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
   320
instance
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
   321
  apply standard
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63886
diff changeset
   322
  unfolding euclidean_representation_sum'
56189
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
   323
  apply (auto simp: less_eq_vec_def inf_vec_def sup_vec_def Inf_vec_def Sup_vec_def inner_axis
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
   324
    Basis_vec_def inner_Basis_inf_left inner_Basis_sup_left inner_Basis_INF_left
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
   325
    inner_Basis_SUP_left eucl_le[where 'a='a] less_le_not_le abs_vec_def abs_inner)
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
   326
  done
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
   327
54780
6fae499e0827 summarized notions related to ordered_euclidean_space and intervals in separate theory
immler
parents:
diff changeset
   328
end
56189
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
   329
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   330
lemma ANR_interval [iff]:
63469
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
   331
    fixes a :: "'a::ordered_euclidean_space"
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
   332
    shows "ANR{a..b}"
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
   333
by (simp add: interval_cbox)
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
   334
69730
0c3dcb3a17f6 tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   335
lemma ENR_interval [iff]:
63469
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
   336
    fixes a :: "'a::ordered_euclidean_space"
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
   337
    shows "ENR{a..b}"
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
   338
  by (auto simp: interval_cbox)
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
   339
56189
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
   340
end
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
   341