| author | wenzelm | 
| Sun, 16 Oct 2022 15:23:07 +0200 | |
| changeset 76316 | 7563367690a1 | 
| parent 71633 | 07bec530f02e | 
| child 77935 | 7f240b0dabd9 | 
| permissions | -rw-r--r-- | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 1 | section \<open>Ordered Euclidean Space\<close> | 
| 69631 | 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 | 
| 71025 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70641diff
changeset | 5 | Convex_Euclidean_Space | 
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
64758diff
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: 
69683diff
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: 
69000diff
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: 
69000diff
changeset | 17 | assumes eucl_Sup: "Sup X = (\<Sum>i\<in>Basis. (SUP x\<in>X. x \<bullet> i) *\<^sub>R i)" | 
| 61945 | 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 | 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 | 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 | 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 | 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 | 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: 
69683diff
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: 
61945diff
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: 
61945diff
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: 
69683diff
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)" | 
| 71633 | 51 | by (simp_all add: eucl_inf eucl_sup inner_sum_left inner_Basis if_distrib | 
| 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: 
69683diff
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: 
69000diff
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: 
69730diff
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: 
69683diff
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: 
69683diff
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: 
69683diff
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 | 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: 
69683diff
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: 
69683diff
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: 
69683diff
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: 
69683diff
changeset | 84 | lemma Sup_eq_maximum_componentwise: | 
| 54781 | 85 | fixes s::"'a set" | 
| 86 | assumes i: "\<And>b. b \<in> Basis \<Longrightarrow> X \<bullet> b = i b \<bullet> b" | |
| 87 | assumes sup: "\<And>b x. b \<in> Basis \<Longrightarrow> x \<in> s \<Longrightarrow> x \<bullet> b \<le> X \<bullet> b" | |
| 88 | assumes i_s: "\<And>b. b \<in> Basis \<Longrightarrow> (i b \<bullet> b) \<in> (\<lambda>x. x \<bullet> b) ` s" | |
| 89 | shows "Sup s = X" | |
| 90 | using assms | |
| 64267 | 91 | unfolding eucl_Sup euclidean_representation_sum | 
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61945diff
changeset | 92 | by (auto intro!: conditionally_complete_lattice_class.cSup_eq_maximum) | 
| 54781 | 93 | |
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 94 | lemma Inf_eq_minimum_componentwise: | 
| 54781 | 95 | assumes i: "\<And>b. b \<in> Basis \<Longrightarrow> X \<bullet> b = i b \<bullet> b" | 
| 96 | assumes sup: "\<And>b x. b \<in> Basis \<Longrightarrow> x \<in> s \<Longrightarrow> X \<bullet> b \<le> x \<bullet> b" | |
| 97 | assumes i_s: "\<And>b. b \<in> Basis \<Longrightarrow> (i b \<bullet> b) \<in> (\<lambda>x. x \<bullet> b) ` s" | |
| 98 | shows "Inf s = X" | |
| 99 | using assms | |
| 64267 | 100 | unfolding eucl_Inf euclidean_representation_sum | 
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61945diff
changeset | 101 | by (auto intro!: conditionally_complete_lattice_class.cInf_eq_minimum) | 
| 54781 | 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: 
69683diff
changeset | 105 | proposition compact_attains_Inf_componentwise: | 
| 54781 | 106 | fixes b::"'a::ordered_euclidean_space" | 
| 107 |   assumes "b \<in> Basis" assumes "X \<noteq> {}" "compact X"
 | |
| 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: 
69683diff
changeset | 109 | proof atomize_elim | 
| 54781 | 110 | let ?proj = "(\<lambda>x. x \<bullet> b) ` X" | 
| 111 |   from assms have "compact ?proj" "?proj \<noteq> {}"
 | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56190diff
changeset | 112 | by (auto intro!: compact_continuous_image continuous_intros) | 
| 54781 | 113 | from compact_attains_inf[OF this] | 
| 114 | obtain s x | |
| 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" | |
| 116 | and x: "x \<in> X" "s = x \<bullet> b" "\<And>y. y \<in> X \<Longrightarrow> x \<bullet> b \<le> y \<bullet> b" | |
| 117 | by auto | |
| 118 | hence "Inf ?proj = x \<bullet> b" | |
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61945diff
changeset | 119 | by (auto intro!: conditionally_complete_lattice_class.cInf_eq_minimum) | 
| 54781 | 120 | hence "x \<bullet> b = Inf X \<bullet> b" | 
| 71633 | 121 | by (auto simp: eucl_Inf inner_sum_left inner_Basis if_distrib \<open>b \<in> Basis\<close> | 
| 56166 | 122 | cong: if_cong) | 
| 54781 | 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 | 
| 124 | qed | |
| 125 | ||
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 126 | proposition | 
| 54781 | 127 | compact_attains_Sup_componentwise: | 
| 128 | fixes b::"'a::ordered_euclidean_space" | |
| 129 |   assumes "b \<in> Basis" assumes "X \<noteq> {}" "compact X"
 | |
| 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: 
69683diff
changeset | 131 | proof atomize_elim | 
| 54781 | 132 | let ?proj = "(\<lambda>x. x \<bullet> b) ` X" | 
| 133 |   from assms have "compact ?proj" "?proj \<noteq> {}"
 | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56190diff
changeset | 134 | by (auto intro!: compact_continuous_image continuous_intros) | 
| 54781 | 135 | from compact_attains_sup[OF this] | 
| 136 | obtain s x | |
| 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" | |
| 138 | and x: "x \<in> X" "s = x \<bullet> b" "\<And>y. y \<in> X \<Longrightarrow> y \<bullet> b \<le> x \<bullet> b" | |
| 139 | by auto | |
| 140 | hence "Sup ?proj = x \<bullet> b" | |
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61945diff
changeset | 141 | by (auto intro!: cSup_eq_maximum) | 
| 54781 | 142 | hence "x \<bullet> b = Sup X \<bullet> b" | 
| 71633 | 143 | by (auto simp: eucl_Sup[where 'a='a] inner_sum_left inner_Basis if_distrib \<open>b \<in> Basis\<close> | 
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61945diff
changeset | 144 | cong: if_cong) | 
| 54781 | 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 | 
| 146 | qed | |
| 147 | ||
| 70378 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70196diff
changeset | 148 | lemma tendsto_sup[tendsto_intros]: | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70196diff
changeset | 149 | fixes X :: "'a \<Rightarrow> 'b::ordered_euclidean_space" | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70196diff
changeset | 150 | assumes "(X \<longlongrightarrow> x) net" "(Y \<longlongrightarrow> y) net" | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70196diff
changeset | 151 | shows "((\<lambda>i. sup (X i) (Y i)) \<longlongrightarrow> sup x y) net" | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70196diff
changeset | 152 | unfolding sup_max eucl_sup by (intro assms tendsto_intros) | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70196diff
changeset | 153 | |
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70196diff
changeset | 154 | lemma tendsto_inf[tendsto_intros]: | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70196diff
changeset | 155 | fixes X :: "'a \<Rightarrow> 'b::ordered_euclidean_space" | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70196diff
changeset | 156 | assumes "(X \<longlongrightarrow> x) net" "(Y \<longlongrightarrow> y) net" | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70196diff
changeset | 157 | shows "((\<lambda>i. inf (X i) (Y i)) \<longlongrightarrow> inf x y) net" | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70196diff
changeset | 158 | unfolding inf_min eucl_inf by (intro assms tendsto_intros) | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70196diff
changeset | 159 | |
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70196diff
changeset | 160 | lemma tendsto_componentwise_max: | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70196diff
changeset | 161 | assumes f: "(f \<longlongrightarrow> l) F" and g: "(g \<longlongrightarrow> m) F" | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70196diff
changeset | 162 | shows "((\<lambda>x. (\<Sum>i\<in>Basis. max (f x \<bullet> i) (g x \<bullet> i) *\<^sub>R i)) \<longlongrightarrow> (\<Sum>i\<in>Basis. max (l \<bullet> i) (m \<bullet> i) *\<^sub>R i)) F" | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70196diff
changeset | 163 | by (intro tendsto_intros assms) | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70196diff
changeset | 164 | |
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70196diff
changeset | 165 | lemma tendsto_componentwise_min: | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70196diff
changeset | 166 | assumes f: "(f \<longlongrightarrow> l) F" and g: "(g \<longlongrightarrow> m) F" | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70196diff
changeset | 167 | shows "((\<lambda>x. (\<Sum>i\<in>Basis. min (f x \<bullet> i) (g x \<bullet> i) *\<^sub>R i)) \<longlongrightarrow> (\<Sum>i\<in>Basis. min (l \<bullet> i) (m \<bullet> i) *\<^sub>R i)) F" | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70196diff
changeset | 168 | by (intro tendsto_intros assms) | 
| 
ebd108578ab1
more new material about analysis
 paulson <lp15@cam.ac.uk> parents: 
70196diff
changeset | 169 | |
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 170 | lemma (in order) atLeastatMost_empty'[simp]: | 
| 69508 | 171 |   "(\<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 | 172 | by (auto) | 
| 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 immler parents: diff
changeset | 173 | |
| 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 immler parents: diff
changeset | 174 | instance real :: ordered_euclidean_space | 
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61945diff
changeset | 175 | by standard auto | 
| 54780 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 immler parents: diff
changeset | 176 | |
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 177 | lemma in_Basis_prod_iff: | 
| 54780 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 immler parents: diff
changeset | 178 | 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 | 179 | 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 | 180 | 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 | 181 | |
| 70136 | 182 | 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 | 183 | begin | 
| 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 immler parents: diff
changeset | 184 | |
| 61945 | 185 | 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 | 186 | |
| 61945 | 187 | instance .. | 
| 188 | ||
| 54780 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 immler parents: diff
changeset | 189 | end | 
| 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 immler parents: diff
changeset | 190 | |
| 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 immler parents: diff
changeset | 191 | instance prod :: (ordered_euclidean_space, ordered_euclidean_space) ordered_euclidean_space | 
| 61169 | 192 | by standard | 
| 64267 | 193 | (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 | 194 | 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 | 195 | 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 | 196 | 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 | 197 | |
| 61808 | 198 | text\<open>Instantiation for intervals on \<open>ordered_euclidean_space\<close>\<close> | 
| 56188 | 199 | |
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 200 | proposition | 
| 61076 | 201 | fixes a :: "'a::ordered_euclidean_space" | 
| 56188 | 202 |   shows cbox_interval: "cbox a b = {a..b}"
 | 
| 203 |     and interval_cbox: "{a..b} = cbox a b"
 | |
| 204 |     and eucl_le_atMost: "{x. \<forall>i\<in>Basis. x \<bullet> i <= a \<bullet> i} = {..a}"
 | |
| 205 |     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: 
69683diff
changeset | 206 | by (auto simp: eucl_le[where 'a='a] eucl_less_def box_def cbox_def) | 
| 56188 | 207 | |
| 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: 
70136diff
changeset | 208 | 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: 
70136diff
changeset | 209 | 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: 
70136diff
changeset | 210 | 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: 
70136diff
changeset | 211 | 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: 
70136diff
changeset | 212 | 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: 
70136diff
changeset | 213 | |
| 
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: 
70136diff
changeset | 214 | 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: 
70136diff
changeset | 215 | 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: 
70136diff
changeset | 216 | 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: 
70136diff
changeset | 217 | 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: 
70136diff
changeset | 218 | 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: 
70136diff
changeset | 219 | |
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 220 | lemma closed_eucl_atLeastAtMost[simp, intro]: | 
| 61076 | 221 | fixes a :: "'a::ordered_euclidean_space" | 
| 56188 | 222 |   shows "closed {a..b}"
 | 
| 223 | by (simp add: cbox_interval[symmetric] closed_cbox) | |
| 224 | ||
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 225 | lemma closed_eucl_atMost[simp, intro]: | 
| 61076 | 226 | fixes a :: "'a::ordered_euclidean_space" | 
| 56188 | 227 |   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: 
66453diff
changeset | 228 | by (simp add: closed_interval_left eucl_le_atMost[symmetric]) | 
| 56188 | 229 | |
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 230 | lemma closed_eucl_atLeast[simp, intro]: | 
| 61076 | 231 | fixes a :: "'a::ordered_euclidean_space" | 
| 56188 | 232 |   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: 
66453diff
changeset | 233 | by (simp add: closed_interval_right eucl_le_atLeast[symmetric]) | 
| 56188 | 234 | |
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 235 | lemma bounded_closed_interval [simp]: | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 236 | fixes a :: "'a::ordered_euclidean_space" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 237 |   shows "bounded {a .. b}"
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 238 | using bounded_cbox[of a b] | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 239 | by (metis interval_cbox) | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 240 | |
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 241 | lemma convex_closed_interval [simp]: | 
| 56190 | 242 | fixes a :: "'a::ordered_euclidean_space" | 
| 243 |   shows "convex {a .. b}"
 | |
| 244 | using convex_box[of a b] | |
| 245 | by (metis interval_cbox) | |
| 246 | ||
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 247 | lemma image_smult_interval:"(\<lambda>x. m *\<^sub>R (x::_::ordered_euclidean_space)) ` {a .. b} =
 | 
| 56188 | 248 |   (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})"
 | 
| 249 | using image_smult_cbox[of m a b] | |
| 250 | by (simp add: cbox_interval) | |
| 54780 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 immler parents: diff
changeset | 251 | |
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 252 | 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: 
69939diff
changeset | 253 | 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: 
69939diff
changeset | 254 |   shows is_interval_ic: "is_interval {..a}"
 | 
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
66827diff
changeset | 255 |     and is_interval_ci: "is_interval {a..}"
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
66827diff
changeset | 256 |     and is_interval_cc: "is_interval {b..a}"
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
66827diff
changeset | 257 | 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: 
66827diff
changeset | 258 | |
| 69000 | 259 | lemma connected_interval [simp]: | 
| 260 | fixes a b::"'a::ordered_euclidean_space" | |
| 261 |   shows "connected {a..b}"
 | |
| 262 | using is_interval_cc is_interval_connected by blast | |
| 263 | ||
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 264 | lemma compact_interval [simp]: | 
| 56190 | 265 | fixes a b::"'a::ordered_euclidean_space" | 
| 266 |   shows "compact {a .. b}"
 | |
| 267 | by (metis compact_cbox interval_cbox) | |
| 268 | ||
| 54780 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 immler parents: diff
changeset | 269 | no_notation | 
| 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 immler parents: diff
changeset | 270 | eucl_less (infix "<e" 50) | 
| 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 immler parents: diff
changeset | 271 | |
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 272 | lemma One_nonneg: "0 \<le> (\<Sum>Basis::'a::ordered_euclidean_space)" | 
| 64267 | 273 | by (auto intro: sum_nonneg) | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 274 | |
| 69730 
0c3dcb3a17f6
tagged 5 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 275 | lemma | 
| 56190 | 276 | fixes a b::"'a::ordered_euclidean_space" | 
| 277 | shows bdd_above_cbox[intro, simp]: "bdd_above (cbox a b)" | |
| 278 | and bdd_below_cbox[intro, simp]: "bdd_below (cbox a b)" | |
| 279 | and bdd_above_box[intro, simp]: "bdd_above (box a b)" | |
| 280 | and bdd_below_box[intro, simp]: "bdd_below (box a b)" | |
| 281 | unfolding atomize_conj | |
| 282 | by (metis bdd_above_Icc bdd_above_mono bdd_below_Icc bdd_below_mono bounded_box | |
| 68120 | 283 | bounded_subset_cbox_symmetric interval_cbox) | 
| 56190 | 284 | |
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 285 | instantiation vec :: (ordered_euclidean_space, finite) ordered_euclidean_space | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 286 | begin | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 287 | |
| 70136 | 288 | definition\<^marker>\<open>tag important\<close> "inf x y = (\<chi> i. inf (x $ i) (y $ i))" | 
| 289 | definition\<^marker>\<open>tag important\<close> "sup x y = (\<chi> i. sup (x $ i) (y $ i))" | |
| 290 | definition\<^marker>\<open>tag important\<close> "Inf X = (\<chi> i. (INF x\<in>X. x $ i))" | |
| 291 | definition\<^marker>\<open>tag important\<close> "Sup X = (\<chi> i. (SUP x\<in>X. x $ i))" | |
| 292 | 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: 
56188diff
changeset | 293 | |
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 294 | instance | 
| 61169 | 295 | apply standard | 
| 64267 | 296 | unfolding euclidean_representation_sum' | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 297 | 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: 
56188diff
changeset | 298 | 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: 
56188diff
changeset | 299 | 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: 
56188diff
changeset | 300 | done | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 301 | |
| 54780 
6fae499e0827
summarized notions related to ordered_euclidean_space and intervals in separate theory
 immler parents: diff
changeset | 302 | end | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 303 | |
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 304 | end | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 305 |