| author | nipkow | 
| Thu, 27 Dec 2018 19:48:28 +0100 | |
| changeset 69508 | 2a4c8a2a3f8e | 
| parent 69260 | 0a9688695a1b | 
| child 69631 | 6c3e6038e74c | 
| permissions | -rw-r--r-- | 
| 
54780
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
1  | 
theory Ordered_Euclidean_Space  | 
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
2  | 
imports  | 
| 
56189
 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 
immler 
parents: 
56188 
diff
changeset
 | 
3  | 
Cartesian_Euclidean_Space  | 
| 
66453
 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 
wenzelm 
parents: 
64758 
diff
changeset
 | 
4  | 
"HOL-Library.Product_Order"  | 
| 
54780
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
5  | 
begin  | 
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
6  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
7  | 
subsection%important \<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
 | 
8  | 
|
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
9  | 
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
 | 
10  | 
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
 | 
11  | 
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
 | 
12  | 
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
 | 
13  | 
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
 | 
14  | 
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
 | 
15  | 
assumes eucl_Sup: "Sup X = (\<Sum>i\<in>Basis. (SUP x\<in>X. x \<bullet> i) *\<^sub>R i)"  | 
| 61945 | 16  | 
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
 | 
17  | 
begin  | 
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
18  | 
|
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
19  | 
subclass order  | 
| 61169 | 20  | 
by standard  | 
| 
54780
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
21  | 
(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
 | 
22  | 
|
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
23  | 
subclass ordered_ab_group_add_abs  | 
| 61169 | 24  | 
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
 | 
25  | 
|
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
26  | 
subclass ordered_real_vector  | 
| 61169 | 27  | 
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
 | 
28  | 
|
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
29  | 
subclass lattice  | 
| 61169 | 30  | 
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
 | 
31  | 
|
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
32  | 
subclass distrib_lattice  | 
| 61169 | 33  | 
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
 | 
34  | 
|
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
35  | 
subclass conditionally_complete_lattice  | 
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
36  | 
proof%unimportant  | 
| 
54780
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
37  | 
fix z::'a and X::"'a set"  | 
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
38  | 
  assume "X \<noteq> {}"
 | 
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
39  | 
  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
 | 
40  | 
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
 | 
41  | 
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
 | 
42  | 
intro!: cInf_greatest cSup_least)  | 
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
43  | 
qed (force intro!: cInf_lower cSup_upper  | 
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
44  | 
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
 | 
45  | 
eucl_Inf eucl_Sup eucl_le)+  | 
| 
54780
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
46  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
47  | 
lemma%unimportant 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
 | 
48  | 
and inner_Basis_sup_left: "i \<in> Basis \<Longrightarrow> sup x y \<bullet> i = sup (x \<bullet> i) (y \<bullet> i)"  | 
| 64267 | 49  | 
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
 | 
50  | 
cong: if_cong)  | 
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
51  | 
|
| 
69260
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
69000 
diff
changeset
 | 
52  | 
lemma%unimportant inner_Basis_INF_left: "i \<in> Basis \<Longrightarrow> (INF x\<in>X. f x) \<bullet> i = (INF x\<in>X. f x \<bullet> i)"  | 
| 
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
69000 
diff
changeset
 | 
53  | 
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)"  | 
| 56166 | 54  | 
using eucl_Sup [of "f ` X"] eucl_Inf [of "f ` X"] by (simp_all add: comp_def)  | 
| 
54780
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
55  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
56  | 
lemma%unimportant 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
 | 
57  | 
by (auto simp: eucl_abs)  | 
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
58  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
59  | 
lemma%unimportant  | 
| 
54780
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
60  | 
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
 | 
61  | 
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
 | 
62  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
63  | 
lemma%unimportant interval_inner_leI:  | 
| 
54780
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
64  | 
  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
 | 
65  | 
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
 | 
66  | 
using assms  | 
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
67  | 
unfolding euclidean_inner[of a i] euclidean_inner[of x i] euclidean_inner[of b i]  | 
| 64267 | 68  | 
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
 | 
69  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
70  | 
lemma%unimportant inner_nonneg_nonneg:  | 
| 
54780
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
71  | 
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
 | 
72  | 
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
 | 
73  | 
by auto  | 
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
74  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
75  | 
lemma%unimportant inner_Basis_mono:  | 
| 
54780
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
76  | 
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
 | 
77  | 
by (simp add: eucl_le)  | 
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
78  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
79  | 
lemma%unimportant 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
 | 
80  | 
by (auto simp: eucl_le inner_Basis)  | 
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
81  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
82  | 
lemma%unimportant Sup_eq_maximum_componentwise:  | 
| 54781 | 83  | 
fixes s::"'a set"  | 
84  | 
assumes i: "\<And>b. b \<in> Basis \<Longrightarrow> X \<bullet> b = i b \<bullet> b"  | 
|
85  | 
assumes sup: "\<And>b x. b \<in> Basis \<Longrightarrow> x \<in> s \<Longrightarrow> x \<bullet> b \<le> X \<bullet> b"  | 
|
86  | 
assumes i_s: "\<And>b. b \<in> Basis \<Longrightarrow> (i b \<bullet> b) \<in> (\<lambda>x. x \<bullet> b) ` s"  | 
|
87  | 
shows "Sup s = X"  | 
|
88  | 
using assms  | 
|
| 64267 | 89  | 
unfolding eucl_Sup euclidean_representation_sum  | 
| 
62343
 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 
haftmann 
parents: 
61945 
diff
changeset
 | 
90  | 
by (auto intro!: conditionally_complete_lattice_class.cSup_eq_maximum)  | 
| 54781 | 91  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
92  | 
lemma%unimportant Inf_eq_minimum_componentwise:  | 
| 54781 | 93  | 
assumes i: "\<And>b. b \<in> Basis \<Longrightarrow> X \<bullet> b = i b \<bullet> b"  | 
94  | 
assumes sup: "\<And>b x. b \<in> Basis \<Longrightarrow> x \<in> s \<Longrightarrow> X \<bullet> b \<le> x \<bullet> b"  | 
|
95  | 
assumes i_s: "\<And>b. b \<in> Basis \<Longrightarrow> (i b \<bullet> b) \<in> (\<lambda>x. x \<bullet> b) ` s"  | 
|
96  | 
shows "Inf s = X"  | 
|
97  | 
using assms  | 
|
| 64267 | 98  | 
unfolding eucl_Inf euclidean_representation_sum  | 
| 
62343
 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 
haftmann 
parents: 
61945 
diff
changeset
 | 
99  | 
by (auto intro!: conditionally_complete_lattice_class.cInf_eq_minimum)  | 
| 54781 | 100  | 
|
| 
54780
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
101  | 
end  | 
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
102  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
103  | 
lemma%important  | 
| 54781 | 104  | 
compact_attains_Inf_componentwise:  | 
105  | 
fixes b::"'a::ordered_euclidean_space"  | 
|
106  | 
  assumes "b \<in> Basis" assumes "X \<noteq> {}" "compact X"
 | 
|
107  | 
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"  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
108  | 
proof%unimportant atomize_elim  | 
| 54781 | 109  | 
let ?proj = "(\<lambda>x. x \<bullet> b) ` X"  | 
110  | 
  from assms have "compact ?proj" "?proj \<noteq> {}"
 | 
|
| 
56371
 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 
hoelzl 
parents: 
56190 
diff
changeset
 | 
111  | 
by (auto intro!: compact_continuous_image continuous_intros)  | 
| 54781 | 112  | 
from compact_attains_inf[OF this]  | 
113  | 
obtain s x  | 
|
114  | 
where s: "s\<in>(\<lambda>x. x \<bullet> b) ` X" "\<And>t. t\<in>(\<lambda>x. x \<bullet> b) ` X \<Longrightarrow> s \<le> t"  | 
|
115  | 
and x: "x \<in> X" "s = x \<bullet> b" "\<And>y. y \<in> X \<Longrightarrow> x \<bullet> b \<le> y \<bullet> b"  | 
|
116  | 
by auto  | 
|
117  | 
hence "Inf ?proj = x \<bullet> b"  | 
|
| 
62343
 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 
haftmann 
parents: 
61945 
diff
changeset
 | 
118  | 
by (auto intro!: conditionally_complete_lattice_class.cInf_eq_minimum)  | 
| 54781 | 119  | 
hence "x \<bullet> b = Inf X \<bullet> b"  | 
| 64267 | 120  | 
by (auto simp: eucl_Inf inner_sum_left inner_Basis if_distrib \<open>b \<in> Basis\<close> sum.delta  | 
| 56166 | 121  | 
cong: if_cong)  | 
| 54781 | 122  | 
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  | 
123  | 
qed  | 
|
124  | 
||
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
125  | 
lemma%important  | 
| 54781 | 126  | 
compact_attains_Sup_componentwise:  | 
127  | 
fixes b::"'a::ordered_euclidean_space"  | 
|
128  | 
  assumes "b \<in> Basis" assumes "X \<noteq> {}" "compact X"
 | 
|
129  | 
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"  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
130  | 
proof%unimportant atomize_elim  | 
| 54781 | 131  | 
let ?proj = "(\<lambda>x. x \<bullet> b) ` X"  | 
132  | 
  from assms have "compact ?proj" "?proj \<noteq> {}"
 | 
|
| 
56371
 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 
hoelzl 
parents: 
56190 
diff
changeset
 | 
133  | 
by (auto intro!: compact_continuous_image continuous_intros)  | 
| 54781 | 134  | 
from compact_attains_sup[OF this]  | 
135  | 
obtain s x  | 
|
136  | 
where s: "s\<in>(\<lambda>x. x \<bullet> b) ` X" "\<And>t. t\<in>(\<lambda>x. x \<bullet> b) ` X \<Longrightarrow> t \<le> s"  | 
|
137  | 
and x: "x \<in> X" "s = x \<bullet> b" "\<And>y. y \<in> X \<Longrightarrow> y \<bullet> b \<le> x \<bullet> b"  | 
|
138  | 
by auto  | 
|
139  | 
hence "Sup ?proj = x \<bullet> b"  | 
|
| 
62343
 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 
haftmann 
parents: 
61945 
diff
changeset
 | 
140  | 
by (auto intro!: cSup_eq_maximum)  | 
| 54781 | 141  | 
hence "x \<bullet> b = Sup X \<bullet> b"  | 
| 64267 | 142  | 
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
 | 
143  | 
cong: if_cong)  | 
| 54781 | 144  | 
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  | 
145  | 
qed  | 
|
146  | 
||
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
147  | 
lemma%unimportant (in order) atLeastatMost_empty'[simp]:  | 
| 69508 | 148  | 
  "(\<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
 | 
149  | 
by (auto)  | 
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
150  | 
|
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
151  | 
instance real :: ordered_euclidean_space  | 
| 
62343
 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 
haftmann 
parents: 
61945 
diff
changeset
 | 
152  | 
by standard auto  | 
| 
54780
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
153  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
154  | 
lemma%unimportant in_Basis_prod_iff:  | 
| 
54780
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
155  | 
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
 | 
156  | 
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
 | 
157  | 
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
 | 
158  | 
|
| 61945 | 159  | 
instantiation prod :: (abs, abs) abs  | 
| 
54780
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
160  | 
begin  | 
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
161  | 
|
| 61945 | 162  | 
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
 | 
163  | 
|
| 61945 | 164  | 
instance ..  | 
165  | 
||
| 
54780
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
166  | 
end  | 
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
167  | 
|
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
168  | 
instance prod :: (ordered_euclidean_space, ordered_euclidean_space) ordered_euclidean_space  | 
| 61169 | 169  | 
by standard  | 
| 64267 | 170  | 
(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
 | 
171  | 
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
 | 
172  | 
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
 | 
173  | 
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
 | 
174  | 
|
| 61808 | 175  | 
text\<open>Instantiation for intervals on \<open>ordered_euclidean_space\<close>\<close>  | 
| 56188 | 176  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
177  | 
lemma%important  | 
| 61076 | 178  | 
fixes a :: "'a::ordered_euclidean_space"  | 
| 56188 | 179  | 
  shows cbox_interval: "cbox a b = {a..b}"
 | 
180  | 
    and interval_cbox: "{a..b} = cbox a b"
 | 
|
181  | 
    and eucl_le_atMost: "{x. \<forall>i\<in>Basis. x \<bullet> i <= a \<bullet> i} = {..a}"
 | 
|
182  | 
    and eucl_le_atLeast: "{x. \<forall>i\<in>Basis. a \<bullet> i <= x \<bullet> i} = {a..}"
 | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
183  | 
by%unimportant (auto simp: eucl_le[where 'a='a] eucl_less_def box_def cbox_def)  | 
| 56188 | 184  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
185  | 
lemma%unimportant 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
 | 
186  | 
fixes a b :: real  | 
| 
 
349c639e593c
more new theorems on real^1, matrices, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
67685 
diff
changeset
 | 
187  | 
  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
 | 
188  | 
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
 | 
189  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
190  | 
lemma%unimportant closed_eucl_atLeastAtMost[simp, intro]:  | 
| 61076 | 191  | 
fixes a :: "'a::ordered_euclidean_space"  | 
| 56188 | 192  | 
  shows "closed {a..b}"
 | 
193  | 
by (simp add: cbox_interval[symmetric] closed_cbox)  | 
|
194  | 
||
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
195  | 
lemma%unimportant closed_eucl_atMost[simp, intro]:  | 
| 61076 | 196  | 
fixes a :: "'a::ordered_euclidean_space"  | 
| 56188 | 197  | 
  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
 | 
198  | 
by (simp add: closed_interval_left eucl_le_atMost[symmetric])  | 
| 56188 | 199  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
200  | 
lemma%unimportant closed_eucl_atLeast[simp, intro]:  | 
| 61076 | 201  | 
fixes a :: "'a::ordered_euclidean_space"  | 
| 56188 | 202  | 
  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
 | 
203  | 
by (simp add: closed_interval_right eucl_le_atLeast[symmetric])  | 
| 56188 | 204  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
205  | 
lemma%unimportant bounded_closed_interval [simp]:  | 
| 
56189
 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 
immler 
parents: 
56188 
diff
changeset
 | 
206  | 
fixes a :: "'a::ordered_euclidean_space"  | 
| 
 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 
immler 
parents: 
56188 
diff
changeset
 | 
207  | 
  shows "bounded {a .. b}"
 | 
| 
 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 
immler 
parents: 
56188 
diff
changeset
 | 
208  | 
using bounded_cbox[of a b]  | 
| 
 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 
immler 
parents: 
56188 
diff
changeset
 | 
209  | 
by (metis interval_cbox)  | 
| 
 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 
immler 
parents: 
56188 
diff
changeset
 | 
210  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
211  | 
lemma%unimportant convex_closed_interval [simp]:  | 
| 56190 | 212  | 
fixes a :: "'a::ordered_euclidean_space"  | 
213  | 
  shows "convex {a .. b}"
 | 
|
214  | 
using convex_box[of a b]  | 
|
215  | 
by (metis interval_cbox)  | 
|
216  | 
||
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
217  | 
lemma%unimportant image_smult_interval:"(\<lambda>x. m *\<^sub>R (x::_::ordered_euclidean_space)) ` {a .. b} =
 | 
| 56188 | 218  | 
  (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})"
 | 
219  | 
using image_smult_cbox[of m a b]  | 
|
220  | 
by (simp add: cbox_interval)  | 
|
| 
54780
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
221  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
222  | 
lemma%unimportant [simp]:  | 
| 
67685
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
66827 
diff
changeset
 | 
223  | 
fixes a b::"'a::ordered_euclidean_space" and r s::real  | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
66827 
diff
changeset
 | 
224  | 
  shows is_interval_io: "is_interval {..<r}"
 | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
66827 
diff
changeset
 | 
225  | 
    and is_interval_ic: "is_interval {..a}"
 | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
66827 
diff
changeset
 | 
226  | 
    and is_interval_oi: "is_interval {r<..}"
 | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
66827 
diff
changeset
 | 
227  | 
    and is_interval_ci: "is_interval {a..}"
 | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
66827 
diff
changeset
 | 
228  | 
    and is_interval_oo: "is_interval {r<..<s}"
 | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
66827 
diff
changeset
 | 
229  | 
    and is_interval_oc: "is_interval {r<..s}"
 | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
66827 
diff
changeset
 | 
230  | 
    and is_interval_co: "is_interval {r..<s}"
 | 
| 
 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 
immler 
parents: 
66827 
diff
changeset
 | 
231  | 
    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
 | 
232  | 
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
 | 
233  | 
|
| 69000 | 234  | 
lemma connected_interval [simp]:  | 
235  | 
fixes a b::"'a::ordered_euclidean_space"  | 
|
236  | 
  shows "connected {a..b}"
 | 
|
237  | 
using is_interval_cc is_interval_connected by blast  | 
|
238  | 
||
239  | 
lemma path_connected_interval [simp]:  | 
|
240  | 
fixes a b::"'a::ordered_euclidean_space"  | 
|
241  | 
  shows "path_connected {a..b}"
 | 
|
242  | 
using is_interval_cc is_interval_path_connected by blast  | 
|
243  | 
||
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
244  | 
lemma%unimportant 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
 | 
245  | 
by (auto simp: real_atLeastGreaterThan_eq)  | 
| 
56189
 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 
immler 
parents: 
56188 
diff
changeset
 | 
246  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
247  | 
lemma%unimportant compact_interval [simp]:  | 
| 56190 | 248  | 
fixes a b::"'a::ordered_euclidean_space"  | 
249  | 
  shows "compact {a .. b}"
 | 
|
250  | 
by (metis compact_cbox interval_cbox)  | 
|
251  | 
||
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
252  | 
lemma%unimportant 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
 | 
253  | 
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
 | 
254  | 
  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
 | 
255  | 
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
 | 
256  | 
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
 | 
257  | 
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
 | 
258  | 
done  | 
| 
 
d21dab28b3f9
New results about paths, segments, etc. The notion of simply_connected.
 
paulson <lp15@cam.ac.uk> 
parents: 
62376 
diff
changeset
 | 
259  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
260  | 
lemma%unimportant 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
 | 
261  | 
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
 | 
262  | 
assumes "a<b" and "c<d"  | 
| 68239 | 263  | 
  shows "{a..b} homeomorphic {c..d}"
 | 
264  | 
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
 | 
265  | 
|
| 
54780
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
266  | 
no_notation  | 
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
267  | 
eucl_less (infix "<e" 50)  | 
| 
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
268  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
269  | 
lemma%unimportant One_nonneg: "0 \<le> (\<Sum>Basis::'a::ordered_euclidean_space)"  | 
| 64267 | 270  | 
by (auto intro: sum_nonneg)  | 
| 
56189
 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 
immler 
parents: 
56188 
diff
changeset
 | 
271  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
272  | 
lemma%unimportant  | 
| 56190 | 273  | 
fixes a b::"'a::ordered_euclidean_space"  | 
274  | 
shows bdd_above_cbox[intro, simp]: "bdd_above (cbox a b)"  | 
|
275  | 
and bdd_below_cbox[intro, simp]: "bdd_below (cbox a b)"  | 
|
276  | 
and bdd_above_box[intro, simp]: "bdd_above (box a b)"  | 
|
277  | 
and bdd_below_box[intro, simp]: "bdd_below (box a b)"  | 
|
278  | 
unfolding atomize_conj  | 
|
279  | 
by (metis bdd_above_Icc bdd_above_mono bdd_below_Icc bdd_below_mono bounded_box  | 
|
| 68120 | 280  | 
bounded_subset_cbox_symmetric interval_cbox)  | 
| 56190 | 281  | 
|
| 
56189
 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 
immler 
parents: 
56188 
diff
changeset
 | 
282  | 
instantiation vec :: (ordered_euclidean_space, finite) ordered_euclidean_space  | 
| 
 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 
immler 
parents: 
56188 
diff
changeset
 | 
283  | 
begin  | 
| 
 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 
immler 
parents: 
56188 
diff
changeset
 | 
284  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
285  | 
definition%important "inf x y = (\<chi> i. inf (x $ i) (y $ i))"  | 
| 
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
286  | 
definition%important "sup x y = (\<chi> i. sup (x $ i) (y $ i))"  | 
| 
69260
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
69000 
diff
changeset
 | 
287  | 
definition%important "Inf X = (\<chi> i. (INF x\<in>X. x $ i))"  | 
| 
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
69000 
diff
changeset
 | 
288  | 
definition%important "Sup X = (\<chi> i. (SUP x\<in>X. x $ i))"  | 
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
289  | 
definition%important "\<bar>x\<bar> = (\<chi> i. \<bar>x $ i\<bar>)"  | 
| 
56189
 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 
immler 
parents: 
56188 
diff
changeset
 | 
290  | 
|
| 
 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 
immler 
parents: 
56188 
diff
changeset
 | 
291  | 
instance  | 
| 61169 | 292  | 
apply standard  | 
| 64267 | 293  | 
unfolding euclidean_representation_sum'  | 
| 
56189
 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 
immler 
parents: 
56188 
diff
changeset
 | 
294  | 
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
 | 
295  | 
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
 | 
296  | 
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
 | 
297  | 
done  | 
| 
 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 
immler 
parents: 
56188 
diff
changeset
 | 
298  | 
|
| 
54780
 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 
immler 
parents:  
diff
changeset
 | 
299  | 
end  | 
| 
56189
 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 
immler 
parents: 
56188 
diff
changeset
 | 
300  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
301  | 
lemma%unimportant ANR_interval [iff]:  | 
| 
63469
 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
62620 
diff
changeset
 | 
302  | 
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
 | 
303  | 
    shows "ANR{a..b}"
 | 
| 
 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
62620 
diff
changeset
 | 
304  | 
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
 | 
305  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68239 
diff
changeset
 | 
306  | 
lemma%unimportant ENR_interval [iff]:  | 
| 
63469
 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
62620 
diff
changeset
 | 
307  | 
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
 | 
308  | 
    shows "ENR{a..b}"
 | 
| 
 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
62620 
diff
changeset
 | 
309  | 
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
 | 
310  | 
|
| 
56189
 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 
immler 
parents: 
56188 
diff
changeset
 | 
311  | 
end  | 
| 
 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 
immler 
parents: 
56188 
diff
changeset
 | 
312  |