| author | wenzelm | 
| Fri, 14 Mar 2025 23:03:58 +0100 | |
| changeset 82276 | d22e9c5b5dc6 | 
| parent 82080 | 0aa2d1c132b2 | 
| child 82470 | 785615e37846 | 
| permissions | -rw-r--r-- | 
| 51524 | 1 | (* Title: HOL/Real_Vector_Spaces.thy | 
| 27552 
15cf4ed9c2a1
re-removed subclass relation real_field < field_char_0: coregularity violation in NSA/HyperDef
 haftmann parents: 
27515diff
changeset | 2 | Author: Brian Huffman | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 3 | Author: Johannes Hölzl | 
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 4 | *) | 
| 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 5 | |
| 60758 | 6 | section \<open>Vector Spaces and Algebras over the Reals\<close> | 
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 7 | |
| 70630 | 8 | theory Real_Vector_Spaces | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 9 | imports Real Topological_Spaces Vector_Spaces | 
| 70630 | 10 | begin | 
| 28029 
4c55cdec4ce7
simplify definition of vector_space locale (use axclasses instead of inheriting from field and ab_group_add classes)
 huffman parents: 
28009diff
changeset | 11 | |
| 60758 | 12 | subsection \<open>Real vector spaces\<close> | 
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 13 | |
| 29608 | 14 | class scaleR = | 
| 80932 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 wenzelm parents: 
79945diff
changeset | 15 | fixes scaleR :: "real \<Rightarrow> 'a \<Rightarrow> 'a" (infixr \<open>*\<^sub>R\<close> 75) | 
| 24748 | 16 | begin | 
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 17 | |
| 80932 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 wenzelm parents: 
79945diff
changeset | 18 | abbreviation divideR :: "'a \<Rightarrow> real \<Rightarrow> 'a" (infixl \<open>'/\<^sub>R\<close> 70) | 
| 70630 | 19 | where "x /\<^sub>R r \<equiv> inverse r *\<^sub>R x" | 
| 24748 | 20 | |
| 21 | end | |
| 22 | ||
| 24588 | 23 | class real_vector = scaleR + ab_group_add + | 
| 70630 | 24 | assumes scaleR_add_right: "a *\<^sub>R (x + y) = a *\<^sub>R x + a *\<^sub>R y" | 
| 25 | and scaleR_add_left: "(a + b) *\<^sub>R x = a *\<^sub>R x + b *\<^sub>R x" | |
| 26 | and scaleR_scaleR: "a *\<^sub>R b *\<^sub>R x = (a * b) *\<^sub>R x" | |
| 27 | and scaleR_one: "1 *\<^sub>R x = x" | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 28 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 29 | class real_algebra = real_vector + ring + | 
| 70630 | 30 | assumes mult_scaleR_left [simp]: "a *\<^sub>R x * y = a *\<^sub>R (x * y)" | 
| 31 | and mult_scaleR_right [simp]: "x * a *\<^sub>R y = a *\<^sub>R (x * y)" | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 32 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 33 | class real_algebra_1 = real_algebra + ring_1 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 34 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 35 | class real_div_algebra = real_algebra_1 + division_ring | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 36 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 37 | class real_field = real_div_algebra + field | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 38 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 39 | instantiation real :: real_field | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 40 | begin | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 41 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 42 | definition real_scaleR_def [simp]: "scaleR a x = a * x" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 43 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 44 | instance | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 45 | by standard (simp_all add: algebra_simps) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 46 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 47 | end | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 48 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 49 | locale linear = Vector_Spaces.linear "scaleR::_\<Rightarrow>_\<Rightarrow>'a::real_vector" "scaleR::_\<Rightarrow>_\<Rightarrow>'b::real_vector" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 50 | begin | 
| 70630 | 51 | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 52 | lemmas scaleR = scale | 
| 70630 | 53 | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 54 | end | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 55 | |
| 70630 | 56 | global_interpretation real_vector?: vector_space "scaleR :: real \<Rightarrow> 'a \<Rightarrow> 'a :: real_vector" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68721diff
changeset | 57 | rewrites "Vector_Spaces.linear (*\<^sub>R) (*\<^sub>R) = linear" | 
| 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68721diff
changeset | 58 | and "Vector_Spaces.linear (*) (*\<^sub>R) = linear" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 59 | defines dependent_raw_def: dependent = real_vector.dependent | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 60 | and representation_raw_def: representation = real_vector.representation | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 61 | and subspace_raw_def: subspace = real_vector.subspace | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 62 | and span_raw_def: span = real_vector.span | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 63 | and extend_basis_raw_def: extend_basis = real_vector.extend_basis | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 64 | and dim_raw_def: dim = real_vector.dim | 
| 71720 | 65 | proof unfold_locales | 
| 66 | show "Vector_Spaces.linear (*\<^sub>R) (*\<^sub>R) = linear" "Vector_Spaces.linear (*) (*\<^sub>R) = linear" | |
| 67 | by (force simp: linear_def real_scaleR_def[abs_def])+ | |
| 68 | qed (use scaleR_add_right scaleR_add_left scaleR_scaleR scaleR_one in auto) | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 69 | |
| 68397 | 70 | hide_const (open)\<comment> \<open>locale constants\<close> | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 71 | real_vector.dependent | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 72 | real_vector.independent | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 73 | real_vector.representation | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 74 | real_vector.subspace | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 75 | real_vector.span | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 76 | real_vector.extend_basis | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 77 | real_vector.dim | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 78 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 79 | abbreviation "independent x \<equiv> \<not> dependent x" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 80 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 81 | global_interpretation real_vector?: vector_space_pair "scaleR::_\<Rightarrow>_\<Rightarrow>'a::real_vector" "scaleR::_\<Rightarrow>_\<Rightarrow>'b::real_vector" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68721diff
changeset | 82 | rewrites "Vector_Spaces.linear (*\<^sub>R) (*\<^sub>R) = linear" | 
| 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68721diff
changeset | 83 | and "Vector_Spaces.linear (*) (*\<^sub>R) = linear" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 84 | defines construct_raw_def: construct = real_vector.construct | 
| 71720 | 85 | proof unfold_locales | 
| 86 | show "Vector_Spaces.linear (*) (*\<^sub>R) = linear" | |
| 87 | unfolding linear_def real_scaleR_def by auto | |
| 88 | qed (auto simp: linear_def) | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 89 | |
| 68397 | 90 | hide_const (open)\<comment> \<open>locale constants\<close> | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 91 | real_vector.construct | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 92 | |
| 68594 | 93 | lemma linear_compose: "linear f \<Longrightarrow> linear g \<Longrightarrow> linear (g \<circ> f)" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 94 | unfolding linear_def by (rule Vector_Spaces.linear_compose) | 
| 28009 
e93b121074fb
move real_vector class proofs into vector_space and group_hom locales
 huffman parents: 
27553diff
changeset | 95 | |
| 60758 | 96 | text \<open>Recover original theorem names\<close> | 
| 28009 
e93b121074fb
move real_vector class proofs into vector_space and group_hom locales
 huffman parents: 
27553diff
changeset | 97 | |
| 
e93b121074fb
move real_vector class proofs into vector_space and group_hom locales
 huffman parents: 
27553diff
changeset | 98 | lemmas scaleR_left_commute = real_vector.scale_left_commute | 
| 
e93b121074fb
move real_vector class proofs into vector_space and group_hom locales
 huffman parents: 
27553diff
changeset | 99 | lemmas scaleR_zero_left = real_vector.scale_zero_left | 
| 
e93b121074fb
move real_vector class proofs into vector_space and group_hom locales
 huffman parents: 
27553diff
changeset | 100 | lemmas scaleR_minus_left = real_vector.scale_minus_left | 
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44127diff
changeset | 101 | lemmas scaleR_diff_left = real_vector.scale_left_diff_distrib | 
| 64267 | 102 | lemmas scaleR_sum_left = real_vector.scale_sum_left | 
| 28009 
e93b121074fb
move real_vector class proofs into vector_space and group_hom locales
 huffman parents: 
27553diff
changeset | 103 | lemmas scaleR_zero_right = real_vector.scale_zero_right | 
| 
e93b121074fb
move real_vector class proofs into vector_space and group_hom locales
 huffman parents: 
27553diff
changeset | 104 | lemmas scaleR_minus_right = real_vector.scale_minus_right | 
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44127diff
changeset | 105 | lemmas scaleR_diff_right = real_vector.scale_right_diff_distrib | 
| 64267 | 106 | lemmas scaleR_sum_right = real_vector.scale_sum_right | 
| 28009 
e93b121074fb
move real_vector class proofs into vector_space and group_hom locales
 huffman parents: 
27553diff
changeset | 107 | lemmas scaleR_eq_0_iff = real_vector.scale_eq_0_iff | 
| 
e93b121074fb
move real_vector class proofs into vector_space and group_hom locales
 huffman parents: 
27553diff
changeset | 108 | lemmas scaleR_left_imp_eq = real_vector.scale_left_imp_eq | 
| 
e93b121074fb
move real_vector class proofs into vector_space and group_hom locales
 huffman parents: 
27553diff
changeset | 109 | lemmas scaleR_right_imp_eq = real_vector.scale_right_imp_eq | 
| 
e93b121074fb
move real_vector class proofs into vector_space and group_hom locales
 huffman parents: 
27553diff
changeset | 110 | lemmas scaleR_cancel_left = real_vector.scale_cancel_left | 
| 
e93b121074fb
move real_vector class proofs into vector_space and group_hom locales
 huffman parents: 
27553diff
changeset | 111 | lemmas scaleR_cancel_right = real_vector.scale_cancel_right | 
| 
e93b121074fb
move real_vector class proofs into vector_space and group_hom locales
 huffman parents: 
27553diff
changeset | 112 | |
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 113 | lemma [field_simps]: | 
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 114 | "c \<noteq> 0 \<Longrightarrow> a = b /\<^sub>R c \<longleftrightarrow> c *\<^sub>R a = b" | 
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 115 | "c \<noteq> 0 \<Longrightarrow> b /\<^sub>R c = a \<longleftrightarrow> b = c *\<^sub>R a" | 
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 116 | "c \<noteq> 0 \<Longrightarrow> a + b /\<^sub>R c = (c *\<^sub>R a + b) /\<^sub>R c" | 
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 117 | "c \<noteq> 0 \<Longrightarrow> a /\<^sub>R c + b = (a + c *\<^sub>R b) /\<^sub>R c" | 
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 118 | "c \<noteq> 0 \<Longrightarrow> a - b /\<^sub>R c = (c *\<^sub>R a - b) /\<^sub>R c" | 
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 119 | "c \<noteq> 0 \<Longrightarrow> a /\<^sub>R c - b = (a - c *\<^sub>R b) /\<^sub>R c" | 
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 120 | "c \<noteq> 0 \<Longrightarrow> - (a /\<^sub>R c) + b = (- a + c *\<^sub>R b) /\<^sub>R c" | 
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 121 | "c \<noteq> 0 \<Longrightarrow> - (a /\<^sub>R c) - b = (- a - c *\<^sub>R b) /\<^sub>R c" | 
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 122 | for a b :: "'a :: real_vector" | 
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 123 | by (auto simp add: scaleR_add_right scaleR_add_left scaleR_diff_right scaleR_diff_left) | 
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 124 | |
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 125 | |
| 60758 | 126 | text \<open>Legacy names\<close> | 
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44127diff
changeset | 127 | |
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44127diff
changeset | 128 | lemmas scaleR_left_distrib = scaleR_add_left | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44127diff
changeset | 129 | lemmas scaleR_right_distrib = scaleR_add_right | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44127diff
changeset | 130 | lemmas scaleR_left_diff_distrib = scaleR_diff_left | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44127diff
changeset | 131 | lemmas scaleR_right_diff_distrib = scaleR_diff_right | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44127diff
changeset | 132 | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 133 | lemmas linear_injective_0 = linear_inj_iff_eq_0 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 134 | and linear_injective_on_subspace_0 = linear_inj_on_iff_eq_0 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 135 | and linear_cmul = linear_scale | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 136 | and linear_scaleR = linear_scale_self | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 137 | and subspace_mul = subspace_scale | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 138 | and span_linear_image = linear_span_image | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 139 | and span_0 = span_zero | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 140 | and span_mul = span_scale | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 141 | and injective_scaleR = injective_scale | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 142 | |
| 63545 | 143 | lemma scaleR_minus1_left [simp]: "scaleR (-1) x = - x" | 
| 144 | for x :: "'a::real_vector" | |
| 78656 
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
 paulson <lp15@cam.ac.uk> parents: 
77221diff
changeset | 145 | by simp | 
| 62101 | 146 | |
| 64788 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64272diff
changeset | 147 | lemma scaleR_2: | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64272diff
changeset | 148 | fixes x :: "'a::real_vector" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64272diff
changeset | 149 | shows "scaleR 2 x = x + x" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64272diff
changeset | 150 | unfolding one_add_one [symmetric] scaleR_left_distrib by simp | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64272diff
changeset | 151 | |
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64272diff
changeset | 152 | lemma scaleR_half_double [simp]: | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64272diff
changeset | 153 | fixes a :: "'a::real_vector" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64272diff
changeset | 154 | shows "(1 / 2) *\<^sub>R (a + a) = a" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64272diff
changeset | 155 | proof - | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64272diff
changeset | 156 | have "\<And>r. r *\<^sub>R (a + a) = (r * 2) *\<^sub>R a" | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64272diff
changeset | 157 | by (metis scaleR_2 scaleR_scaleR) | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64272diff
changeset | 158 | then show ?thesis | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64272diff
changeset | 159 | by simp | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64272diff
changeset | 160 | qed | 
| 
19f3d4af7a7d
New material about path connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
64272diff
changeset | 161 | |
| 77138 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 paulson <lp15@cam.ac.uk> parents: 
76055diff
changeset | 162 | lemma shift_zero_ident [simp]: | 
| 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 paulson <lp15@cam.ac.uk> parents: 
76055diff
changeset | 163 | fixes f :: "'a \<Rightarrow> 'b::real_vector" | 
| 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 paulson <lp15@cam.ac.uk> parents: 
76055diff
changeset | 164 | shows "(+)0 \<circ> f = f" | 
| 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 paulson <lp15@cam.ac.uk> parents: 
76055diff
changeset | 165 | by force | 
| 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 paulson <lp15@cam.ac.uk> parents: 
76055diff
changeset | 166 | |
| 70019 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69700diff
changeset | 167 | lemma linear_scale_real: | 
| 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69700diff
changeset | 168 | fixes r::real shows "linear f \<Longrightarrow> f (r * b) = r * f b" | 
| 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69700diff
changeset | 169 | using linear_scale by fastforce | 
| 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69700diff
changeset | 170 | |
| 63545 | 171 | interpretation scaleR_left: additive "(\<lambda>a. scaleR a x :: 'a::real_vector)" | 
| 172 | by standard (rule scaleR_left_distrib) | |
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 173 | |
| 63545 | 174 | interpretation scaleR_right: additive "(\<lambda>x. scaleR a x :: 'a::real_vector)" | 
| 175 | by standard (rule scaleR_right_distrib) | |
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 176 | |
| 20584 
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
 huffman parents: 
20560diff
changeset | 177 | lemma nonzero_inverse_scaleR_distrib: | 
| 63545 | 178 | "a \<noteq> 0 \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> inverse (scaleR a x) = scaleR (inverse a) (inverse x)" | 
| 179 | for x :: "'a::real_div_algebra" | |
| 180 | by (rule inverse_unique) simp | |
| 20584 
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
 huffman parents: 
20560diff
changeset | 181 | |
| 63545 | 182 | lemma inverse_scaleR_distrib: "inverse (scaleR a x) = scaleR (inverse a) (inverse x)" | 
| 183 |   for x :: "'a::{real_div_algebra,division_ring}"
 | |
| 68594 | 184 | by (metis inverse_zero nonzero_inverse_scaleR_distrib scale_eq_0_iff) | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 185 | |
| 68397 | 186 | lemmas sum_constant_scaleR = real_vector.sum_constant_scale\<comment> \<open>legacy name\<close> | 
| 63545 | 187 | |
| 63927 
0efb482beb84
vector_add_divide_simps now a "named theorems" bundle
 paulson <lp15@cam.ac.uk> parents: 
63915diff
changeset | 188 | named_theorems vector_add_divide_simps "to simplify sums of scaled vectors" | 
| 
0efb482beb84
vector_add_divide_simps now a "named theorems" bundle
 paulson <lp15@cam.ac.uk> parents: 
63915diff
changeset | 189 | |
| 
0efb482beb84
vector_add_divide_simps now a "named theorems" bundle
 paulson <lp15@cam.ac.uk> parents: 
63915diff
changeset | 190 | lemma [vector_add_divide_simps]: | 
| 63545 | 191 | "v + (b / z) *\<^sub>R w = (if z = 0 then v else (z *\<^sub>R v + b *\<^sub>R w) /\<^sub>R z)" | 
| 192 | "a *\<^sub>R v + (b / z) *\<^sub>R w = (if z = 0 then a *\<^sub>R v else ((a * z) *\<^sub>R v + b *\<^sub>R w) /\<^sub>R z)" | |
| 193 | "(a / z) *\<^sub>R v + w = (if z = 0 then w else (a *\<^sub>R v + z *\<^sub>R w) /\<^sub>R z)" | |
| 194 | "(a / z) *\<^sub>R v + b *\<^sub>R w = (if z = 0 then b *\<^sub>R w else (a *\<^sub>R v + (b * z) *\<^sub>R w) /\<^sub>R z)" | |
| 195 | "v - (b / z) *\<^sub>R w = (if z = 0 then v else (z *\<^sub>R v - b *\<^sub>R w) /\<^sub>R z)" | |
| 196 | "a *\<^sub>R v - (b / z) *\<^sub>R w = (if z = 0 then a *\<^sub>R v else ((a * z) *\<^sub>R v - b *\<^sub>R w) /\<^sub>R z)" | |
| 197 | "(a / z) *\<^sub>R v - w = (if z = 0 then -w else (a *\<^sub>R v - z *\<^sub>R w) /\<^sub>R z)" | |
| 198 | "(a / z) *\<^sub>R v - b *\<^sub>R w = (if z = 0 then -b *\<^sub>R w else (a *\<^sub>R v - (b * z) *\<^sub>R w) /\<^sub>R z)" | |
| 199 | for v :: "'a :: real_vector" | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 200 | by (simp_all add: divide_inverse_commute scaleR_add_right scaleR_diff_right) | 
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63040diff
changeset | 201 | |
| 63927 
0efb482beb84
vector_add_divide_simps now a "named theorems" bundle
 paulson <lp15@cam.ac.uk> parents: 
63915diff
changeset | 202 | |
| 
0efb482beb84
vector_add_divide_simps now a "named theorems" bundle
 paulson <lp15@cam.ac.uk> parents: 
63915diff
changeset | 203 | lemma eq_vector_fraction_iff [vector_add_divide_simps]: | 
| 
0efb482beb84
vector_add_divide_simps now a "named theorems" bundle
 paulson <lp15@cam.ac.uk> parents: 
63915diff
changeset | 204 | fixes x :: "'a :: real_vector" | 
| 
0efb482beb84
vector_add_divide_simps now a "named theorems" bundle
 paulson <lp15@cam.ac.uk> parents: 
63915diff
changeset | 205 | shows "(x = (u / v) *\<^sub>R a) \<longleftrightarrow> (if v=0 then x = 0 else v *\<^sub>R x = u *\<^sub>R a)" | 
| 
0efb482beb84
vector_add_divide_simps now a "named theorems" bundle
 paulson <lp15@cam.ac.uk> parents: 
63915diff
changeset | 206 | by auto (metis (no_types) divide_eq_1_iff divide_inverse_commute scaleR_one scaleR_scaleR) | 
| 
0efb482beb84
vector_add_divide_simps now a "named theorems" bundle
 paulson <lp15@cam.ac.uk> parents: 
63915diff
changeset | 207 | |
| 
0efb482beb84
vector_add_divide_simps now a "named theorems" bundle
 paulson <lp15@cam.ac.uk> parents: 
63915diff
changeset | 208 | lemma vector_fraction_eq_iff [vector_add_divide_simps]: | 
| 
0efb482beb84
vector_add_divide_simps now a "named theorems" bundle
 paulson <lp15@cam.ac.uk> parents: 
63915diff
changeset | 209 | fixes x :: "'a :: real_vector" | 
| 
0efb482beb84
vector_add_divide_simps now a "named theorems" bundle
 paulson <lp15@cam.ac.uk> parents: 
63915diff
changeset | 210 | shows "((u / v) *\<^sub>R a = x) \<longleftrightarrow> (if v=0 then x = 0 else u *\<^sub>R a = v *\<^sub>R x)" | 
| 
0efb482beb84
vector_add_divide_simps now a "named theorems" bundle
 paulson <lp15@cam.ac.uk> parents: 
63915diff
changeset | 211 | by (metis eq_vector_fraction_iff) | 
| 
0efb482beb84
vector_add_divide_simps now a "named theorems" bundle
 paulson <lp15@cam.ac.uk> parents: 
63915diff
changeset | 212 | |
| 60800 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 213 | lemma real_vector_affinity_eq: | 
| 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 214 | fixes x :: "'a :: real_vector" | 
| 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 215 | assumes m0: "m \<noteq> 0" | 
| 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 216 | shows "m *\<^sub>R x + c = y \<longleftrightarrow> x = inverse m *\<^sub>R y - (inverse m *\<^sub>R c)" | 
| 63545 | 217 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 60800 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 218 | proof | 
| 63545 | 219 | assume ?lhs | 
| 220 | then have "m *\<^sub>R x = y - c" by (simp add: field_simps) | |
| 221 | then have "inverse m *\<^sub>R (m *\<^sub>R x) = inverse m *\<^sub>R (y - c)" by simp | |
| 60800 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 222 | then show "x = inverse m *\<^sub>R y - (inverse m *\<^sub>R c)" | 
| 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 223 | using m0 | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 224 | by (simp add: scaleR_diff_right) | 
| 60800 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 225 | next | 
| 63545 | 226 | assume ?rhs | 
| 227 | with m0 show "m *\<^sub>R x + c = y" | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 228 | by (simp add: scaleR_diff_right) | 
| 60800 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 229 | qed | 
| 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 230 | |
| 63545 | 231 | lemma real_vector_eq_affinity: "m \<noteq> 0 \<Longrightarrow> y = m *\<^sub>R x + c \<longleftrightarrow> inverse m *\<^sub>R y - (inverse m *\<^sub>R c) = x" | 
| 232 | for x :: "'a::real_vector" | |
| 60800 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 233 | using real_vector_affinity_eq[where m=m and x=x and y=y and c=c] | 
| 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 234 | by metis | 
| 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 235 | |
| 63545 | 236 | lemma scaleR_eq_iff [simp]: "b + u *\<^sub>R a = a + u *\<^sub>R b \<longleftrightarrow> a = b \<or> u = 1" | 
| 237 | for a :: "'a::real_vector" | |
| 238 | proof (cases "u = 1") | |
| 239 | case True | |
| 240 | then show ?thesis by auto | |
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62623diff
changeset | 241 | next | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62623diff
changeset | 242 | case False | 
| 63545 | 243 | have "a = b" if "b + u *\<^sub>R a = a + u *\<^sub>R b" | 
| 244 | proof - | |
| 245 | from that have "(u - 1) *\<^sub>R a = (u - 1) *\<^sub>R b" | |
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62623diff
changeset | 246 | by (simp add: algebra_simps) | 
| 63545 | 247 | with False show ?thesis | 
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62623diff
changeset | 248 | by auto | 
| 63545 | 249 | qed | 
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62623diff
changeset | 250 | then show ?thesis by auto | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62623diff
changeset | 251 | qed | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62623diff
changeset | 252 | |
| 63545 | 253 | lemma scaleR_collapse [simp]: "(1 - u) *\<^sub>R a + u *\<^sub>R a = a" | 
| 254 | for a :: "'a::real_vector" | |
| 255 | by (simp add: algebra_simps) | |
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62623diff
changeset | 256 | |
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 257 | |
| 63545 | 258 | subsection \<open>Embedding of the Reals into any \<open>real_algebra_1\<close>: \<open>of_real\<close>\<close> | 
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 259 | |
| 63545 | 260 | definition of_real :: "real \<Rightarrow> 'a::real_algebra_1" | 
| 261 | where "of_real r = scaleR r 1" | |
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 262 | |
| 21809 
4b93e949ac33
remove uses of scaleR infix syntax; add lemma Reals_number_of
 huffman parents: 
21404diff
changeset | 263 | lemma scaleR_conv_of_real: "scaleR r x = of_real r * x" | 
| 63545 | 264 | by (simp add: of_real_def) | 
| 20763 | 265 | |
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 266 | lemma of_real_0 [simp]: "of_real 0 = 0" | 
| 63545 | 267 | by (simp add: of_real_def) | 
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 268 | |
| 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 269 | lemma of_real_1 [simp]: "of_real 1 = 1" | 
| 63545 | 270 | by (simp add: of_real_def) | 
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 271 | |
| 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 272 | lemma of_real_add [simp]: "of_real (x + y) = of_real x + of_real y" | 
| 63545 | 273 | by (simp add: of_real_def scaleR_left_distrib) | 
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 274 | |
| 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 275 | lemma of_real_minus [simp]: "of_real (- x) = - of_real x" | 
| 63545 | 276 | by (simp add: of_real_def) | 
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 277 | |
| 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 278 | lemma of_real_diff [simp]: "of_real (x - y) = of_real x - of_real y" | 
| 63545 | 279 | by (simp add: of_real_def scaleR_left_diff_distrib) | 
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 280 | |
| 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 281 | lemma of_real_mult [simp]: "of_real (x * y) = of_real x * of_real y" | 
| 71544 | 282 | by (simp add: of_real_def) | 
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 283 | |
| 64267 | 284 | lemma of_real_sum[simp]: "of_real (sum f s) = (\<Sum>x\<in>s. of_real (f x))" | 
| 56889 
48a745e1bde7
avoid the Complex constructor, use the more natural Re/Im view; moved csqrt to Complex.
 hoelzl parents: 
56479diff
changeset | 285 | by (induct s rule: infinite_finite_induct) auto | 
| 
48a745e1bde7
avoid the Complex constructor, use the more natural Re/Im view; moved csqrt to Complex.
 hoelzl parents: 
56479diff
changeset | 286 | |
| 64272 | 287 | lemma of_real_prod[simp]: "of_real (prod f s) = (\<Prod>x\<in>s. of_real (f x))" | 
| 56889 
48a745e1bde7
avoid the Complex constructor, use the more natural Re/Im view; moved csqrt to Complex.
 hoelzl parents: 
56479diff
changeset | 288 | by (induct s rule: infinite_finite_induct) auto | 
| 
48a745e1bde7
avoid the Complex constructor, use the more natural Re/Im view; moved csqrt to Complex.
 hoelzl parents: 
56479diff
changeset | 289 | |
| 82080 
0aa2d1c132b2
A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
 paulson <lp15@cam.ac.uk> parents: 
80932diff
changeset | 290 | lemma sum_list_of_real: "sum_list (map of_real xs) = of_real (sum_list xs)" | 
| 
0aa2d1c132b2
A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
 paulson <lp15@cam.ac.uk> parents: 
80932diff
changeset | 291 | by (induction xs) auto | 
| 
0aa2d1c132b2
A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
 paulson <lp15@cam.ac.uk> parents: 
80932diff
changeset | 292 | |
| 20584 
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
 huffman parents: 
20560diff
changeset | 293 | lemma nonzero_of_real_inverse: | 
| 63545 | 294 | "x \<noteq> 0 \<Longrightarrow> of_real (inverse x) = inverse (of_real x :: 'a::real_div_algebra)" | 
| 295 | by (simp add: of_real_def nonzero_inverse_scaleR_distrib) | |
| 20584 
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
 huffman parents: 
20560diff
changeset | 296 | |
| 
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
 huffman parents: 
20560diff
changeset | 297 | lemma of_real_inverse [simp]: | 
| 63545 | 298 |   "of_real (inverse x) = inverse (of_real x :: 'a::{real_div_algebra,division_ring})"
 | 
| 299 | by (simp add: of_real_def inverse_scaleR_distrib) | |
| 20584 
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
 huffman parents: 
20560diff
changeset | 300 | |
| 
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
 huffman parents: 
20560diff
changeset | 301 | lemma nonzero_of_real_divide: | 
| 63545 | 302 | "y \<noteq> 0 \<Longrightarrow> of_real (x / y) = (of_real x / of_real y :: 'a::real_field)" | 
| 303 | by (simp add: divide_inverse nonzero_of_real_inverse) | |
| 20722 | 304 | |
| 305 | lemma of_real_divide [simp]: | |
| 62131 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62102diff
changeset | 306 | "of_real (x / y) = (of_real x / of_real y :: 'a::real_div_algebra)" | 
| 63545 | 307 | by (simp add: divide_inverse) | 
| 20584 
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
 huffman parents: 
20560diff
changeset | 308 | |
| 20722 | 309 | lemma of_real_power [simp]: | 
| 31017 | 310 |   "of_real (x ^ n) = (of_real x :: 'a::{real_algebra_1}) ^ n"
 | 
| 63545 | 311 | by (induct n) simp_all | 
| 20722 | 312 | |
| 71837 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 313 | lemma of_real_power_int [simp]: | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 314 |   "of_real (power_int x n) = power_int (of_real x :: 'a :: {real_div_algebra,division_ring}) n"
 | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 315 | by (auto simp: power_int_def) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 316 | |
| 63545 | 317 | lemma of_real_eq_iff [simp]: "of_real x = of_real y \<longleftrightarrow> x = y" | 
| 318 | by (simp add: of_real_def) | |
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 319 | |
| 63545 | 320 | lemma inj_of_real: "inj of_real" | 
| 38621 
d6cb7e625d75
more concise characterization of of_nat operation and class semiring_char_0
 haftmann parents: 
37887diff
changeset | 321 | by (auto intro: injI) | 
| 
d6cb7e625d75
more concise characterization of of_nat operation and class semiring_char_0
 haftmann parents: 
37887diff
changeset | 322 | |
| 20584 
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
 huffman parents: 
20560diff
changeset | 323 | lemmas of_real_eq_0_iff [simp] = of_real_eq_iff [of _ 0, simplified] | 
| 65578 
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
 paulson <lp15@cam.ac.uk> parents: 
65036diff
changeset | 324 | lemmas of_real_eq_1_iff [simp] = of_real_eq_iff [of _ 1, simplified] | 
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 325 | |
| 67135 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66793diff
changeset | 326 | lemma minus_of_real_eq_of_real_iff [simp]: "-of_real x = of_real y \<longleftrightarrow> -x = y" | 
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66793diff
changeset | 327 | using of_real_eq_iff[of "-x" y] by (simp only: of_real_minus) | 
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66793diff
changeset | 328 | |
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66793diff
changeset | 329 | lemma of_real_eq_minus_of_real_iff [simp]: "of_real x = -of_real y \<longleftrightarrow> x = -y" | 
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66793diff
changeset | 330 | using of_real_eq_iff[of x "-y"] by (simp only: of_real_minus) | 
| 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 Manuel Eberl <eberlm@in.tum.de> parents: 
66793diff
changeset | 331 | |
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 332 | lemma of_real_eq_id [simp]: "of_real = (id :: real \<Rightarrow> real)" | 
| 63545 | 333 | by (rule ext) (simp add: of_real_def) | 
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 334 | |
| 63545 | 335 | text \<open>Collapse nested embeddings.\<close> | 
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 336 | lemma of_real_of_nat_eq [simp]: "of_real (of_nat n) = of_nat n" | 
| 63545 | 337 | by (induct n) auto | 
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 338 | |
| 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 339 | lemma of_real_of_int_eq [simp]: "of_real (of_int z) = of_int z" | 
| 63545 | 340 | by (cases z rule: int_diff_cases) simp | 
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 341 | |
| 60155 
91477b3a2d6b
Tidying. Improved simplification for numerals, esp in exponents.
 paulson <lp15@cam.ac.uk> parents: 
60026diff
changeset | 342 | lemma of_real_numeral [simp]: "of_real (numeral w) = numeral w" | 
| 63545 | 343 | using of_real_of_int_eq [of "numeral w"] by simp | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46868diff
changeset | 344 | |
| 60155 
91477b3a2d6b
Tidying. Improved simplification for numerals, esp in exponents.
 paulson <lp15@cam.ac.uk> parents: 
60026diff
changeset | 345 | lemma of_real_neg_numeral [simp]: "of_real (- numeral w) = - numeral w" | 
| 63545 | 346 | using of_real_of_int_eq [of "- numeral w"] by simp | 
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 347 | |
| 71837 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 348 | lemma numeral_power_int_eq_of_real_cancel_iff [simp]: | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 349 |   "power_int (numeral x) n = (of_real y :: 'a :: {real_div_algebra, division_ring}) \<longleftrightarrow>
 | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 350 | power_int (numeral x) n = y" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 351 | proof - | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 352 | have "power_int (numeral x) n = (of_real (power_int (numeral x) n) :: 'a)" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 353 | by simp | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 354 | also have "\<dots> = of_real y \<longleftrightarrow> power_int (numeral x) n = y" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 355 | by (subst of_real_eq_iff) auto | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 356 | finally show ?thesis . | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 357 | qed | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 358 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 359 | lemma of_real_eq_numeral_power_int_cancel_iff [simp]: | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 360 |   "(of_real y :: 'a :: {real_div_algebra, division_ring}) = power_int (numeral x) n \<longleftrightarrow>
 | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 361 | y = power_int (numeral x) n" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 362 | by (subst (1 2) eq_commute) simp | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 363 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 364 | lemma of_real_eq_of_real_power_int_cancel_iff [simp]: | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 365 |   "power_int (of_real b :: 'a :: {real_div_algebra, division_ring}) w = of_real x \<longleftrightarrow>
 | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 366 | power_int b w = x" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 367 | by (metis of_real_power_int of_real_eq_iff) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 368 | |
| 73109 
783406dd051e
some algebra material for HOL: characteristic of a ring, algebraic integers
 Manuel Eberl <eberlm@in.tum.de> parents: 
71837diff
changeset | 369 | lemma of_real_in_Ints_iff [simp]: "of_real x \<in> \<int> \<longleftrightarrow> x \<in> \<int>" | 
| 
783406dd051e
some algebra material for HOL: characteristic of a ring, algebraic integers
 Manuel Eberl <eberlm@in.tum.de> parents: 
71837diff
changeset | 370 | proof safe | 
| 
783406dd051e
some algebra material for HOL: characteristic of a ring, algebraic integers
 Manuel Eberl <eberlm@in.tum.de> parents: 
71837diff
changeset | 371 | fix x assume "(of_real x :: 'a) \<in> \<int>" | 
| 
783406dd051e
some algebra material for HOL: characteristic of a ring, algebraic integers
 Manuel Eberl <eberlm@in.tum.de> parents: 
71837diff
changeset | 372 | then obtain n where "(of_real x :: 'a) = of_int n" | 
| 
783406dd051e
some algebra material for HOL: characteristic of a ring, algebraic integers
 Manuel Eberl <eberlm@in.tum.de> parents: 
71837diff
changeset | 373 | by (auto simp: Ints_def) | 
| 
783406dd051e
some algebra material for HOL: characteristic of a ring, algebraic integers
 Manuel Eberl <eberlm@in.tum.de> parents: 
71837diff
changeset | 374 | also have "of_int n = of_real (real_of_int n)" | 
| 
783406dd051e
some algebra material for HOL: characteristic of a ring, algebraic integers
 Manuel Eberl <eberlm@in.tum.de> parents: 
71837diff
changeset | 375 | by simp | 
| 
783406dd051e
some algebra material for HOL: characteristic of a ring, algebraic integers
 Manuel Eberl <eberlm@in.tum.de> parents: 
71837diff
changeset | 376 | finally have "x = real_of_int n" | 
| 
783406dd051e
some algebra material for HOL: characteristic of a ring, algebraic integers
 Manuel Eberl <eberlm@in.tum.de> parents: 
71837diff
changeset | 377 | by (subst (asm) of_real_eq_iff) | 
| 
783406dd051e
some algebra material for HOL: characteristic of a ring, algebraic integers
 Manuel Eberl <eberlm@in.tum.de> parents: 
71837diff
changeset | 378 | thus "x \<in> \<int>" | 
| 
783406dd051e
some algebra material for HOL: characteristic of a ring, algebraic integers
 Manuel Eberl <eberlm@in.tum.de> parents: 
71837diff
changeset | 379 | by auto | 
| 
783406dd051e
some algebra material for HOL: characteristic of a ring, algebraic integers
 Manuel Eberl <eberlm@in.tum.de> parents: 
71837diff
changeset | 380 | qed (auto simp: Ints_def) | 
| 
783406dd051e
some algebra material for HOL: characteristic of a ring, algebraic integers
 Manuel Eberl <eberlm@in.tum.de> parents: 
71837diff
changeset | 381 | |
| 
783406dd051e
some algebra material for HOL: characteristic of a ring, algebraic integers
 Manuel Eberl <eberlm@in.tum.de> parents: 
71837diff
changeset | 382 | lemma Ints_of_real [intro]: "x \<in> \<int> \<Longrightarrow> of_real x \<in> \<int>" | 
| 
783406dd051e
some algebra material for HOL: characteristic of a ring, algebraic integers
 Manuel Eberl <eberlm@in.tum.de> parents: 
71837diff
changeset | 383 | by simp | 
| 
783406dd051e
some algebra material for HOL: characteristic of a ring, algebraic integers
 Manuel Eberl <eberlm@in.tum.de> parents: 
71837diff
changeset | 384 | |
| 
783406dd051e
some algebra material for HOL: characteristic of a ring, algebraic integers
 Manuel Eberl <eberlm@in.tum.de> parents: 
71837diff
changeset | 385 | |
| 63545 | 386 | text \<open>Every real algebra has characteristic zero.\<close> | 
| 22912 | 387 | instance real_algebra_1 < ring_char_0 | 
| 388 | proof | |
| 63545 | 389 | from inj_of_real inj_of_nat have "inj (of_real \<circ> of_nat)" | 
| 69700 
7a92cbec7030
new material about summations and powers, along with some tweaks
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 390 | by (rule inj_compose) | 
| 63545 | 391 | then show "inj (of_nat :: nat \<Rightarrow> 'a)" | 
| 392 | by (simp add: comp_def) | |
| 22912 | 393 | qed | 
| 394 | ||
| 63967 
2aa42596edc3
new material on paths, etc. Also rationalisation
 paulson <lp15@cam.ac.uk> parents: 
63927diff
changeset | 395 | lemma fraction_scaleR_times [simp]: | 
| 
2aa42596edc3
new material on paths, etc. Also rationalisation
 paulson <lp15@cam.ac.uk> parents: 
63927diff
changeset | 396 | fixes a :: "'a::real_algebra_1" | 
| 
2aa42596edc3
new material on paths, etc. Also rationalisation
 paulson <lp15@cam.ac.uk> parents: 
63927diff
changeset | 397 | shows "(numeral u / numeral v) *\<^sub>R (numeral w * a) = (numeral u * numeral w / numeral v) *\<^sub>R a" | 
| 
2aa42596edc3
new material on paths, etc. Also rationalisation
 paulson <lp15@cam.ac.uk> parents: 
63927diff
changeset | 398 | by (metis (no_types, lifting) of_real_numeral scaleR_conv_of_real scaleR_scaleR times_divide_eq_left) | 
| 
2aa42596edc3
new material on paths, etc. Also rationalisation
 paulson <lp15@cam.ac.uk> parents: 
63927diff
changeset | 399 | |
| 
2aa42596edc3
new material on paths, etc. Also rationalisation
 paulson <lp15@cam.ac.uk> parents: 
63927diff
changeset | 400 | lemma inverse_scaleR_times [simp]: | 
| 
2aa42596edc3
new material on paths, etc. Also rationalisation
 paulson <lp15@cam.ac.uk> parents: 
63927diff
changeset | 401 | fixes a :: "'a::real_algebra_1" | 
| 
2aa42596edc3
new material on paths, etc. Also rationalisation
 paulson <lp15@cam.ac.uk> parents: 
63927diff
changeset | 402 | shows "(1 / numeral v) *\<^sub>R (numeral w * a) = (numeral w / numeral v) *\<^sub>R a" | 
| 
2aa42596edc3
new material on paths, etc. Also rationalisation
 paulson <lp15@cam.ac.uk> parents: 
63927diff
changeset | 403 | by (metis divide_inverse_commute inverse_eq_divide of_real_numeral scaleR_conv_of_real scaleR_scaleR) | 
| 
2aa42596edc3
new material on paths, etc. Also rationalisation
 paulson <lp15@cam.ac.uk> parents: 
63927diff
changeset | 404 | |
| 
2aa42596edc3
new material on paths, etc. Also rationalisation
 paulson <lp15@cam.ac.uk> parents: 
63927diff
changeset | 405 | lemma scaleR_times [simp]: | 
| 
2aa42596edc3
new material on paths, etc. Also rationalisation
 paulson <lp15@cam.ac.uk> parents: 
63927diff
changeset | 406 | fixes a :: "'a::real_algebra_1" | 
| 
2aa42596edc3
new material on paths, etc. Also rationalisation
 paulson <lp15@cam.ac.uk> parents: 
63927diff
changeset | 407 | shows "(numeral u) *\<^sub>R (numeral w * a) = (numeral u * numeral w) *\<^sub>R a" | 
| 
2aa42596edc3
new material on paths, etc. Also rationalisation
 paulson <lp15@cam.ac.uk> parents: 
63927diff
changeset | 408 | by (simp add: scaleR_conv_of_real) | 
| 
2aa42596edc3
new material on paths, etc. Also rationalisation
 paulson <lp15@cam.ac.uk> parents: 
63927diff
changeset | 409 | |
| 27553 | 410 | instance real_field < field_char_0 .. | 
| 411 | ||
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 412 | |
| 60758 | 413 | subsection \<open>The Set of Real Numbers\<close> | 
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 414 | |
| 80932 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 wenzelm parents: 
79945diff
changeset | 415 | definition Reals :: "'a::real_algebra_1 set" (\<open>\<real>\<close>) | 
| 61070 | 416 | where "\<real> = range of_real" | 
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 417 | |
| 61070 | 418 | lemma Reals_of_real [simp]: "of_real r \<in> \<real>" | 
| 63545 | 419 | by (simp add: Reals_def) | 
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 420 | |
| 61070 | 421 | lemma Reals_of_int [simp]: "of_int z \<in> \<real>" | 
| 63545 | 422 | by (subst of_real_of_int_eq [symmetric], rule Reals_of_real) | 
| 20718 | 423 | |
| 61070 | 424 | lemma Reals_of_nat [simp]: "of_nat n \<in> \<real>" | 
| 63545 | 425 | by (subst of_real_of_nat_eq [symmetric], rule Reals_of_real) | 
| 21809 
4b93e949ac33
remove uses of scaleR infix syntax; add lemma Reals_number_of
 huffman parents: 
21404diff
changeset | 426 | |
| 61070 | 427 | lemma Reals_numeral [simp]: "numeral w \<in> \<real>" | 
| 63545 | 428 | by (subst of_real_numeral [symmetric], rule Reals_of_real) | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46868diff
changeset | 429 | |
| 68594 | 430 | lemma Reals_0 [simp]: "0 \<in> \<real>" and Reals_1 [simp]: "1 \<in> \<real>" | 
| 431 | by (simp_all add: Reals_def) | |
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 432 | |
| 63545 | 433 | lemma Reals_add [simp]: "a \<in> \<real> \<Longrightarrow> b \<in> \<real> \<Longrightarrow> a + b \<in> \<real>" | 
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
73411diff
changeset | 434 | by (metis (no_types, opaque_lifting) Reals_def Reals_of_real imageE of_real_add) | 
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 435 | |
| 61070 | 436 | lemma Reals_minus [simp]: "a \<in> \<real> \<Longrightarrow> - a \<in> \<real>" | 
| 68594 | 437 | by (auto simp: Reals_def) | 
| 20584 
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
 huffman parents: 
20560diff
changeset | 438 | |
| 68499 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 paulson <lp15@cam.ac.uk> parents: 
68465diff
changeset | 439 | lemma Reals_minus_iff [simp]: "- a \<in> \<real> \<longleftrightarrow> a \<in> \<real>" | 
| 71720 | 440 | using Reals_minus by fastforce | 
| 68499 
d4312962161a
Rationalisation of complex transcendentals, esp the Arg function
 paulson <lp15@cam.ac.uk> parents: 
68465diff
changeset | 441 | |
| 63545 | 442 | lemma Reals_diff [simp]: "a \<in> \<real> \<Longrightarrow> b \<in> \<real> \<Longrightarrow> a - b \<in> \<real>" | 
| 68594 | 443 | by (metis Reals_add Reals_minus_iff add_uminus_conv_diff) | 
| 20584 
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
 huffman parents: 
20560diff
changeset | 444 | |
| 63545 | 445 | lemma Reals_mult [simp]: "a \<in> \<real> \<Longrightarrow> b \<in> \<real> \<Longrightarrow> a * b \<in> \<real>" | 
| 68594 | 446 | by (metis (no_types, lifting) Reals_def Reals_of_real imageE of_real_mult) | 
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 447 | |
| 63545 | 448 | lemma nonzero_Reals_inverse: "a \<in> \<real> \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> inverse a \<in> \<real>" | 
| 449 | for a :: "'a::real_div_algebra" | |
| 68594 | 450 | by (metis Reals_def Reals_of_real imageE of_real_inverse) | 
| 20584 
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
 huffman parents: 
20560diff
changeset | 451 | |
| 63545 | 452 | lemma Reals_inverse: "a \<in> \<real> \<Longrightarrow> inverse a \<in> \<real>" | 
| 453 |   for a :: "'a::{real_div_algebra,division_ring}"
 | |
| 68594 | 454 | using nonzero_Reals_inverse by fastforce | 
| 20584 
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
 huffman parents: 
20560diff
changeset | 455 | |
| 63545 | 456 | lemma Reals_inverse_iff [simp]: "inverse x \<in> \<real> \<longleftrightarrow> x \<in> \<real>" | 
| 457 |   for x :: "'a::{real_div_algebra,division_ring}"
 | |
| 458 | by (metis Reals_inverse inverse_inverse_eq) | |
| 55719 
cdddd073bff8
Lemmas about Reals, norm, etc., and cleaner variants of existing ones
 paulson <lp15@cam.ac.uk> parents: 
54890diff
changeset | 459 | |
| 63545 | 460 | lemma nonzero_Reals_divide: "a \<in> \<real> \<Longrightarrow> b \<in> \<real> \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> a / b \<in> \<real>" | 
| 461 | for a b :: "'a::real_field" | |
| 68594 | 462 | by (simp add: divide_inverse) | 
| 20584 
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
 huffman parents: 
20560diff
changeset | 463 | |
| 63545 | 464 | lemma Reals_divide [simp]: "a \<in> \<real> \<Longrightarrow> b \<in> \<real> \<Longrightarrow> a / b \<in> \<real>" | 
| 465 |   for a b :: "'a::{real_field,field}"
 | |
| 68594 | 466 | using nonzero_Reals_divide by fastforce | 
| 20584 
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
 huffman parents: 
20560diff
changeset | 467 | |
| 63545 | 468 | lemma Reals_power [simp]: "a \<in> \<real> \<Longrightarrow> a ^ n \<in> \<real>" | 
| 469 | for a :: "'a::real_algebra_1" | |
| 68594 | 470 | by (metis Reals_def Reals_of_real imageE of_real_power) | 
| 20722 | 471 | |
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 472 | lemma Reals_cases [cases set: Reals]: | 
| 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 473 | assumes "q \<in> \<real>" | 
| 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 474 | obtains (of_real) r where "q = of_real r" | 
| 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 475 | unfolding Reals_def | 
| 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 476 | proof - | 
| 60758 | 477 | from \<open>q \<in> \<real>\<close> have "q \<in> range of_real" unfolding Reals_def . | 
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 478 | then obtain r where "q = of_real r" .. | 
| 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 479 | then show thesis .. | 
| 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 480 | qed | 
| 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 481 | |
| 64267 | 482 | lemma sum_in_Reals [intro,simp]: "(\<And>i. i \<in> s \<Longrightarrow> f i \<in> \<real>) \<Longrightarrow> sum f s \<in> \<real>" | 
| 63915 | 483 | proof (induct s rule: infinite_finite_induct) | 
| 484 | case infinite | |
| 64267 | 485 | then show ?case by (metis Reals_0 sum.infinite) | 
| 63915 | 486 | qed simp_all | 
| 55719 
cdddd073bff8
Lemmas about Reals, norm, etc., and cleaner variants of existing ones
 paulson <lp15@cam.ac.uk> parents: 
54890diff
changeset | 487 | |
| 64272 | 488 | lemma prod_in_Reals [intro,simp]: "(\<And>i. i \<in> s \<Longrightarrow> f i \<in> \<real>) \<Longrightarrow> prod f s \<in> \<real>" | 
| 63915 | 489 | proof (induct s rule: infinite_finite_induct) | 
| 490 | case infinite | |
| 64272 | 491 | then show ?case by (metis Reals_1 prod.infinite) | 
| 63915 | 492 | qed simp_all | 
| 55719 
cdddd073bff8
Lemmas about Reals, norm, etc., and cleaner variants of existing ones
 paulson <lp15@cam.ac.uk> parents: 
54890diff
changeset | 493 | |
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 494 | lemma Reals_induct [case_names of_real, induct set: Reals]: | 
| 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 495 | "q \<in> \<real> \<Longrightarrow> (\<And>r. P (of_real r)) \<Longrightarrow> P q" | 
| 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 496 | by (rule Reals_cases) auto | 
| 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 497 | |
| 63545 | 498 | |
| 60758 | 499 | subsection \<open>Ordered real vector spaces\<close> | 
| 54778 | 500 | |
| 501 | class ordered_real_vector = real_vector + ordered_ab_group_add + | |
| 502 | assumes scaleR_left_mono: "x \<le> y \<Longrightarrow> 0 \<le> a \<Longrightarrow> a *\<^sub>R x \<le> a *\<^sub>R y" | |
| 63545 | 503 | and scaleR_right_mono: "a \<le> b \<Longrightarrow> 0 \<le> x \<Longrightarrow> a *\<^sub>R x \<le> b *\<^sub>R x" | 
| 54778 | 504 | begin | 
| 505 | ||
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 506 | lemma scaleR_mono: | 
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 507 | "a \<le> b \<Longrightarrow> x \<le> y \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> x \<Longrightarrow> a *\<^sub>R x \<le> b *\<^sub>R y" | 
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 508 | by (meson order_trans scaleR_left_mono scaleR_right_mono) | 
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 509 | |
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 510 | lemma scaleR_mono': | 
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 511 | "a \<le> b \<Longrightarrow> c \<le> d \<Longrightarrow> 0 \<le> a \<Longrightarrow> 0 \<le> c \<Longrightarrow> a *\<^sub>R c \<le> b *\<^sub>R d" | 
| 54778 | 512 | by (rule scaleR_mono) (auto intro: order.trans) | 
| 513 | ||
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 514 | lemma pos_le_divideR_eq [field_simps]: | 
| 70630 | 515 | "a \<le> b /\<^sub>R c \<longleftrightarrow> c *\<^sub>R a \<le> b" (is "?P \<longleftrightarrow> ?Q") if "0 < c" | 
| 516 | proof | |
| 517 | assume ?P | |
| 518 | with scaleR_left_mono that have "c *\<^sub>R a \<le> c *\<^sub>R (b /\<^sub>R c)" | |
| 54785 | 519 | by simp | 
| 70630 | 520 | with that show ?Q | 
| 521 | by (simp add: scaleR_one scaleR_scaleR inverse_eq_divide) | |
| 522 | next | |
| 523 | assume ?Q | |
| 524 | with scaleR_left_mono that have "c *\<^sub>R a /\<^sub>R c \<le> b /\<^sub>R c" | |
| 525 | by simp | |
| 526 | with that show ?P | |
| 54785 | 527 | by (simp add: scaleR_one scaleR_scaleR inverse_eq_divide) | 
| 528 | qed | |
| 529 | ||
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 530 | lemma pos_less_divideR_eq [field_simps]: | 
| 70630 | 531 | "a < b /\<^sub>R c \<longleftrightarrow> c *\<^sub>R a < b" if "c > 0" | 
| 532 | using that pos_le_divideR_eq [of c a b] | |
| 533 | by (auto simp add: le_less scaleR_scaleR scaleR_one) | |
| 534 | ||
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 535 | lemma pos_divideR_le_eq [field_simps]: | 
| 70630 | 536 | "b /\<^sub>R c \<le> a \<longleftrightarrow> b \<le> c *\<^sub>R a" if "c > 0" | 
| 537 | using that pos_le_divideR_eq [of "inverse c" b a] by simp | |
| 538 | ||
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 539 | lemma pos_divideR_less_eq [field_simps]: | 
| 70630 | 540 | "b /\<^sub>R c < a \<longleftrightarrow> b < c *\<^sub>R a" if "c > 0" | 
| 541 | using that pos_less_divideR_eq [of "inverse c" b a] by simp | |
| 542 | ||
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 543 | lemma pos_le_minus_divideR_eq [field_simps]: | 
| 70630 | 544 | "a \<le> - (b /\<^sub>R c) \<longleftrightarrow> c *\<^sub>R a \<le> - b" if "c > 0" | 
| 545 | using that by (metis add_minus_cancel diff_0 left_minus minus_minus neg_le_iff_le | |
| 546 | scaleR_add_right uminus_add_conv_diff pos_le_divideR_eq) | |
| 547 | ||
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 548 | lemma pos_less_minus_divideR_eq [field_simps]: | 
| 70630 | 549 | "a < - (b /\<^sub>R c) \<longleftrightarrow> c *\<^sub>R a < - b" if "c > 0" | 
| 550 | using that by (metis le_less less_le_not_le pos_divideR_le_eq | |
| 551 | pos_divideR_less_eq pos_le_minus_divideR_eq) | |
| 552 | ||
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 553 | lemma pos_minus_divideR_le_eq [field_simps]: | 
| 70630 | 554 | "- (b /\<^sub>R c) \<le> a \<longleftrightarrow> - b \<le> c *\<^sub>R a" if "c > 0" | 
| 555 | using that by (metis pos_divideR_le_eq pos_le_minus_divideR_eq that | |
| 556 | inverse_positive_iff_positive le_imp_neg_le minus_minus) | |
| 557 | ||
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 558 | lemma pos_minus_divideR_less_eq [field_simps]: | 
| 70630 | 559 | "- (b /\<^sub>R c) < a \<longleftrightarrow> - b < c *\<^sub>R a" if "c > 0" | 
| 560 | using that by (simp add: less_le_not_le pos_le_minus_divideR_eq pos_minus_divideR_le_eq) | |
| 54785 | 561 | |
| 63545 | 562 | lemma scaleR_image_atLeastAtMost: "c > 0 \<Longrightarrow> scaleR c ` {x..y} = {c *\<^sub>R x..c *\<^sub>R y}"
 | 
| 71720 | 563 | apply (auto intro!: scaleR_left_mono simp: image_iff Bex_def) | 
| 73411 | 564 | using pos_divideR_le_eq [of c] pos_le_divideR_eq [of c] | 
| 565 | apply (meson local.order_eq_iff) | |
| 566 | done | |
| 54785 | 567 | |
| 54778 | 568 | end | 
| 569 | ||
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 570 | lemma neg_le_divideR_eq [field_simps]: | 
| 70630 | 571 | "a \<le> b /\<^sub>R c \<longleftrightarrow> b \<le> c *\<^sub>R a" (is "?P \<longleftrightarrow> ?Q") if "c < 0" | 
| 572 | for a b :: "'a :: ordered_real_vector" | |
| 573 | using that pos_le_divideR_eq [of "- c" a "- b"] by simp | |
| 574 | ||
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 575 | lemma neg_less_divideR_eq [field_simps]: | 
| 70630 | 576 | "a < b /\<^sub>R c \<longleftrightarrow> b < c *\<^sub>R a" if "c < 0" | 
| 577 | for a b :: "'a :: ordered_real_vector" | |
| 578 | using that neg_le_divideR_eq [of c a b] by (auto simp add: le_less) | |
| 579 | ||
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 580 | lemma neg_divideR_le_eq [field_simps]: | 
| 70630 | 581 | "b /\<^sub>R c \<le> a \<longleftrightarrow> c *\<^sub>R a \<le> b" if "c < 0" | 
| 582 | for a b :: "'a :: ordered_real_vector" | |
| 583 | using that pos_divideR_le_eq [of "- c" "- b" a] by simp | |
| 584 | ||
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 585 | lemma neg_divideR_less_eq [field_simps]: | 
| 70630 | 586 | "b /\<^sub>R c < a \<longleftrightarrow> c *\<^sub>R a < b" if "c < 0" | 
| 587 | for a b :: "'a :: ordered_real_vector" | |
| 588 | using that neg_divideR_le_eq [of c b a] by (auto simp add: le_less) | |
| 589 | ||
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 590 | lemma neg_le_minus_divideR_eq [field_simps]: | 
| 70630 | 591 | "a \<le> - (b /\<^sub>R c) \<longleftrightarrow> - b \<le> c *\<^sub>R a" if "c < 0" | 
| 592 | for a b :: "'a :: ordered_real_vector" | |
| 593 | using that pos_le_minus_divideR_eq [of "- c" a "- b"] by (simp add: minus_le_iff) | |
| 594 | ||
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 595 | lemma neg_less_minus_divideR_eq [field_simps]: | 
| 70630 | 596 | "a < - (b /\<^sub>R c) \<longleftrightarrow> - b < c *\<^sub>R a" if "c < 0" | 
| 597 | for a b :: "'a :: ordered_real_vector" | |
| 598 | proof - | |
| 599 | have *: "- b = c *\<^sub>R a \<longleftrightarrow> b = - (c *\<^sub>R a)" | |
| 600 | by (metis add.inverse_inverse) | |
| 601 | from that neg_le_minus_divideR_eq [of c a b] | |
| 602 | show ?thesis by (auto simp add: le_less *) | |
| 603 | qed | |
| 604 | ||
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 605 | lemma neg_minus_divideR_le_eq [field_simps]: | 
| 70630 | 606 | "- (b /\<^sub>R c) \<le> a \<longleftrightarrow> c *\<^sub>R a \<le> - b" if "c < 0" | 
| 607 | for a b :: "'a :: ordered_real_vector" | |
| 608 | using that pos_minus_divideR_le_eq [of "- c" "- b" a] by (simp add: le_minus_iff) | |
| 609 | ||
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70723diff
changeset | 610 | lemma neg_minus_divideR_less_eq [field_simps]: | 
| 70630 | 611 | "- (b /\<^sub>R c) < a \<longleftrightarrow> c *\<^sub>R a < - b" if "c < 0" | 
| 612 | for a b :: "'a :: ordered_real_vector" | |
| 613 | using that by (simp add: less_le_not_le neg_le_minus_divideR_eq neg_minus_divideR_le_eq) | |
| 60303 | 614 | |
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 615 | lemma [field_split_simps]: | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 616 | "a = b /\<^sub>R c \<longleftrightarrow> (if c = 0 then a = 0 else c *\<^sub>R a = b)" | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 617 | "b /\<^sub>R c = a \<longleftrightarrow> (if c = 0 then a = 0 else b = c *\<^sub>R a)" | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 618 | "a + b /\<^sub>R c = (if c = 0 then a else (c *\<^sub>R a + b) /\<^sub>R c)" | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 619 | "a /\<^sub>R c + b = (if c = 0 then b else (a + c *\<^sub>R b) /\<^sub>R c)" | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 620 | "a - b /\<^sub>R c = (if c = 0 then a else (c *\<^sub>R a - b) /\<^sub>R c)" | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 621 | "a /\<^sub>R c - b = (if c = 0 then - b else (a - c *\<^sub>R b) /\<^sub>R c)" | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 622 | "- (a /\<^sub>R c) + b = (if c = 0 then b else (- a + c *\<^sub>R b) /\<^sub>R c)" | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 623 | "- (a /\<^sub>R c) - b = (if c = 0 then - b else (- a - c *\<^sub>R b) /\<^sub>R c)" | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 624 | for a b :: "'a :: real_vector" | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 625 | by (auto simp add: field_simps) | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 626 | |
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 627 | lemma [field_split_simps]: | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 628 | "0 < c \<Longrightarrow> a \<le> b /\<^sub>R c \<longleftrightarrow> (if c > 0 then c *\<^sub>R a \<le> b else if c < 0 then b \<le> c *\<^sub>R a else a \<le> 0)" | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 629 | "0 < c \<Longrightarrow> a < b /\<^sub>R c \<longleftrightarrow> (if c > 0 then c *\<^sub>R a < b else if c < 0 then b < c *\<^sub>R a else a < 0)" | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 630 | "0 < c \<Longrightarrow> b /\<^sub>R c \<le> a \<longleftrightarrow> (if c > 0 then b \<le> c *\<^sub>R a else if c < 0 then c *\<^sub>R a \<le> b else a \<ge> 0)" | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 631 | "0 < c \<Longrightarrow> b /\<^sub>R c < a \<longleftrightarrow> (if c > 0 then b < c *\<^sub>R a else if c < 0 then c *\<^sub>R a < b else a > 0)" | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 632 | "0 < c \<Longrightarrow> a \<le> - (b /\<^sub>R c) \<longleftrightarrow> (if c > 0 then c *\<^sub>R a \<le> - b else if c < 0 then - b \<le> c *\<^sub>R a else a \<le> 0)" | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 633 | "0 < c \<Longrightarrow> a < - (b /\<^sub>R c) \<longleftrightarrow> (if c > 0 then c *\<^sub>R a < - b else if c < 0 then - b < c *\<^sub>R a else a < 0)" | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 634 | "0 < c \<Longrightarrow> - (b /\<^sub>R c) \<le> a \<longleftrightarrow> (if c > 0 then - b \<le> c *\<^sub>R a else if c < 0 then c *\<^sub>R a \<le> - b else a \<ge> 0)" | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 635 | "0 < c \<Longrightarrow> - (b /\<^sub>R c) < a \<longleftrightarrow> (if c > 0 then - b < c *\<^sub>R a else if c < 0 then c *\<^sub>R a < - b else a > 0)" | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 636 | for a b :: "'a :: ordered_real_vector" | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 637 | by (clarsimp intro!: field_simps)+ | 
| 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 638 | |
| 63545 | 639 | lemma scaleR_nonneg_nonneg: "0 \<le> a \<Longrightarrow> 0 \<le> x \<Longrightarrow> 0 \<le> a *\<^sub>R x" | 
| 640 | for x :: "'a::ordered_real_vector" | |
| 641 | using scaleR_left_mono [of 0 x a] by simp | |
| 54778 | 642 | |
| 63545 | 643 | lemma scaleR_nonneg_nonpos: "0 \<le> a \<Longrightarrow> x \<le> 0 \<Longrightarrow> a *\<^sub>R x \<le> 0" | 
| 644 | for x :: "'a::ordered_real_vector" | |
| 54778 | 645 | using scaleR_left_mono [of x 0 a] by simp | 
| 646 | ||
| 63545 | 647 | lemma scaleR_nonpos_nonneg: "a \<le> 0 \<Longrightarrow> 0 \<le> x \<Longrightarrow> a *\<^sub>R x \<le> 0" | 
| 648 | for x :: "'a::ordered_real_vector" | |
| 54778 | 649 | using scaleR_right_mono [of a 0 x] by simp | 
| 650 | ||
| 63545 | 651 | lemma split_scaleR_neg_le: "(0 \<le> a \<and> x \<le> 0) \<or> (a \<le> 0 \<and> 0 \<le> x) \<Longrightarrow> a *\<^sub>R x \<le> 0" | 
| 652 | for x :: "'a::ordered_real_vector" | |
| 68594 | 653 | by (auto simp: scaleR_nonneg_nonpos scaleR_nonpos_nonneg) | 
| 54778 | 654 | |
| 63545 | 655 | lemma le_add_iff1: "a *\<^sub>R e + c \<le> b *\<^sub>R e + d \<longleftrightarrow> (a - b) *\<^sub>R e + c \<le> d" | 
| 656 | for c d e :: "'a::ordered_real_vector" | |
| 54778 | 657 | by (simp add: algebra_simps) | 
| 658 | ||
| 63545 | 659 | lemma le_add_iff2: "a *\<^sub>R e + c \<le> b *\<^sub>R e + d \<longleftrightarrow> c \<le> (b - a) *\<^sub>R e + d" | 
| 660 | for c d e :: "'a::ordered_real_vector" | |
| 54778 | 661 | by (simp add: algebra_simps) | 
| 662 | ||
| 63545 | 663 | lemma scaleR_left_mono_neg: "b \<le> a \<Longrightarrow> c \<le> 0 \<Longrightarrow> c *\<^sub>R a \<le> c *\<^sub>R b" | 
| 664 | for a b :: "'a::ordered_real_vector" | |
| 68669 | 665 | by (drule scaleR_left_mono [of _ _ "- c"], simp_all) | 
| 54778 | 666 | |
| 63545 | 667 | lemma scaleR_right_mono_neg: "b \<le> a \<Longrightarrow> c \<le> 0 \<Longrightarrow> a *\<^sub>R c \<le> b *\<^sub>R c" | 
| 668 | for c :: "'a::ordered_real_vector" | |
| 68669 | 669 | by (drule scaleR_right_mono [of _ _ "- c"], simp_all) | 
| 54778 | 670 | |
| 63545 | 671 | lemma scaleR_nonpos_nonpos: "a \<le> 0 \<Longrightarrow> b \<le> 0 \<Longrightarrow> 0 \<le> a *\<^sub>R b" | 
| 672 | for b :: "'a::ordered_real_vector" | |
| 673 | using scaleR_right_mono_neg [of a 0 b] by simp | |
| 54778 | 674 | |
| 63545 | 675 | lemma split_scaleR_pos_le: "(0 \<le> a \<and> 0 \<le> b) \<or> (a \<le> 0 \<and> b \<le> 0) \<Longrightarrow> 0 \<le> a *\<^sub>R b" | 
| 676 | for b :: "'a::ordered_real_vector" | |
| 68594 | 677 | by (auto simp: scaleR_nonneg_nonneg scaleR_nonpos_nonpos) | 
| 54778 | 678 | |
| 679 | lemma zero_le_scaleR_iff: | |
| 63545 | 680 | fixes b :: "'a::ordered_real_vector" | 
| 681 | shows "0 \<le> a *\<^sub>R b \<longleftrightarrow> 0 < a \<and> 0 \<le> b \<or> a < 0 \<and> b \<le> 0 \<or> a = 0" | |
| 682 | (is "?lhs = ?rhs") | |
| 683 | proof (cases "a = 0") | |
| 684 | case True | |
| 685 | then show ?thesis by simp | |
| 686 | next | |
| 687 | case False | |
| 54778 | 688 | show ?thesis | 
| 689 | proof | |
| 63545 | 690 | assume ?lhs | 
| 691 | from \<open>a \<noteq> 0\<close> consider "a > 0" | "a < 0" by arith | |
| 692 | then show ?rhs | |
| 693 | proof cases | |
| 694 | case 1 | |
| 695 | with \<open>?lhs\<close> have "inverse a *\<^sub>R 0 \<le> inverse a *\<^sub>R (a *\<^sub>R b)" | |
| 54778 | 696 | by (intro scaleR_mono) auto | 
| 63545 | 697 | with 1 show ?thesis | 
| 54778 | 698 | by simp | 
| 63545 | 699 | next | 
| 700 | case 2 | |
| 701 | with \<open>?lhs\<close> have "- inverse a *\<^sub>R 0 \<le> - inverse a *\<^sub>R (a *\<^sub>R b)" | |
| 54778 | 702 | by (intro scaleR_mono) auto | 
| 63545 | 703 | with 2 show ?thesis | 
| 54778 | 704 | by simp | 
| 63545 | 705 | qed | 
| 706 | next | |
| 707 | assume ?rhs | |
| 708 | then show ?lhs | |
| 709 | by (auto simp: not_le \<open>a \<noteq> 0\<close> intro!: split_scaleR_pos_le) | |
| 710 | qed | |
| 711 | qed | |
| 54778 | 712 | |
| 63545 | 713 | lemma scaleR_le_0_iff: "a *\<^sub>R b \<le> 0 \<longleftrightarrow> 0 < a \<and> b \<le> 0 \<or> a < 0 \<and> 0 \<le> b \<or> a = 0" | 
| 714 | for b::"'a::ordered_real_vector" | |
| 54778 | 715 | by (insert zero_le_scaleR_iff [of "-a" b]) force | 
| 716 | ||
| 63545 | 717 | lemma scaleR_le_cancel_left: "c *\<^sub>R a \<le> c *\<^sub>R b \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> a)" | 
| 718 | for b :: "'a::ordered_real_vector" | |
| 68594 | 719 | by (auto simp: neq_iff scaleR_left_mono scaleR_left_mono_neg | 
| 63545 | 720 | dest: scaleR_left_mono[where a="inverse c"] scaleR_left_mono_neg[where c="inverse c"]) | 
| 54778 | 721 | |
| 63545 | 722 | lemma scaleR_le_cancel_left_pos: "0 < c \<Longrightarrow> c *\<^sub>R a \<le> c *\<^sub>R b \<longleftrightarrow> a \<le> b" | 
| 723 | for b :: "'a::ordered_real_vector" | |
| 54778 | 724 | by (auto simp: scaleR_le_cancel_left) | 
| 725 | ||
| 63545 | 726 | lemma scaleR_le_cancel_left_neg: "c < 0 \<Longrightarrow> c *\<^sub>R a \<le> c *\<^sub>R b \<longleftrightarrow> b \<le> a" | 
| 727 | for b :: "'a::ordered_real_vector" | |
| 54778 | 728 | by (auto simp: scaleR_le_cancel_left) | 
| 729 | ||
| 63545 | 730 | lemma scaleR_left_le_one_le: "0 \<le> x \<Longrightarrow> a \<le> 1 \<Longrightarrow> a *\<^sub>R x \<le> x" | 
| 731 | for x :: "'a::ordered_real_vector" and a :: real | |
| 54778 | 732 | using scaleR_right_mono[of a 1 x] by simp | 
| 733 | ||
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 734 | |
| 60758 | 735 | subsection \<open>Real normed vector spaces\<close> | 
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 736 | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 737 | class dist = | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 738 | fixes dist :: "'a \<Rightarrow> 'a \<Rightarrow> real" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 739 | |
| 29608 | 740 | class norm = | 
| 22636 | 741 | fixes norm :: "'a \<Rightarrow> real" | 
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 742 | |
| 24520 | 743 | class sgn_div_norm = scaleR + norm + sgn + | 
| 25062 | 744 | assumes sgn_div_norm: "sgn x = x /\<^sub>R norm x" | 
| 24506 | 745 | |
| 31289 | 746 | class dist_norm = dist + norm + minus + | 
| 747 | assumes dist_norm: "dist x y = norm (x - y)" | |
| 748 | ||
| 62101 | 749 | class uniformity_dist = dist + uniformity + | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69064diff
changeset | 750 |   assumes uniformity_dist: "uniformity = (INF e\<in>{0 <..}. principal {(x, y). dist x y < e})"
 | 
| 62101 | 751 | begin | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 752 | |
| 62101 | 753 | lemma eventually_uniformity_metric: | 
| 754 | "eventually P uniformity \<longleftrightarrow> (\<exists>e>0. \<forall>x y. dist x y < e \<longrightarrow> P (x, y))" | |
| 755 | unfolding uniformity_dist | |
| 756 | by (subst eventually_INF_base) | |
| 757 | (auto simp: eventually_principal subset_eq intro: bexI[of _ "min _ _"]) | |
| 758 | ||
| 759 | end | |
| 760 | ||
| 761 | class real_normed_vector = real_vector + sgn_div_norm + dist_norm + uniformity_dist + open_uniformity + | |
| 51002 
496013a6eb38
remove unnecessary assumption from real_normed_vector
 hoelzl parents: 
50999diff
changeset | 762 | assumes norm_eq_zero [simp]: "norm x = 0 \<longleftrightarrow> x = 0" | 
| 63545 | 763 | and norm_triangle_ineq: "norm (x + y) \<le> norm x + norm y" | 
| 764 | and norm_scaleR [simp]: "norm (scaleR a x) = \<bar>a\<bar> * norm x" | |
| 51002 
496013a6eb38
remove unnecessary assumption from real_normed_vector
 hoelzl parents: 
50999diff
changeset | 765 | begin | 
| 
496013a6eb38
remove unnecessary assumption from real_normed_vector
 hoelzl parents: 
50999diff
changeset | 766 | |
| 
496013a6eb38
remove unnecessary assumption from real_normed_vector
 hoelzl parents: 
50999diff
changeset | 767 | lemma norm_ge_zero [simp]: "0 \<le> norm x" | 
| 
496013a6eb38
remove unnecessary assumption from real_normed_vector
 hoelzl parents: 
50999diff
changeset | 768 | proof - | 
| 60026 
41d81b4a0a21
Restored LIMSEQ_def as legacy binding. [The other changes are whitespace only.]
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 769 | have "0 = norm (x + -1 *\<^sub>R x)" | 
| 51002 
496013a6eb38
remove unnecessary assumption from real_normed_vector
 hoelzl parents: 
50999diff
changeset | 770 | using scaleR_add_left[of 1 "-1" x] norm_scaleR[of 0 x] by (simp add: scaleR_one) | 
| 
496013a6eb38
remove unnecessary assumption from real_normed_vector
 hoelzl parents: 
50999diff
changeset | 771 | also have "\<dots> \<le> norm x + norm (-1 *\<^sub>R x)" by (rule norm_triangle_ineq) | 
| 
496013a6eb38
remove unnecessary assumption from real_normed_vector
 hoelzl parents: 
50999diff
changeset | 772 | finally show ?thesis by simp | 
| 
496013a6eb38
remove unnecessary assumption from real_normed_vector
 hoelzl parents: 
50999diff
changeset | 773 | qed | 
| 
496013a6eb38
remove unnecessary assumption from real_normed_vector
 hoelzl parents: 
50999diff
changeset | 774 | |
| 74007 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 775 | lemma bdd_below_norm_image: "bdd_below (norm ` A)" | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 776 | by (meson bdd_belowI2 norm_ge_zero) | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 777 | |
| 51002 
496013a6eb38
remove unnecessary assumption from real_normed_vector
 hoelzl parents: 
50999diff
changeset | 778 | end | 
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 779 | |
| 24588 | 780 | class real_normed_algebra = real_algebra + real_normed_vector + | 
| 25062 | 781 | assumes norm_mult_ineq: "norm (x * y) \<le> norm x * norm y" | 
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 782 | |
| 24588 | 783 | class real_normed_algebra_1 = real_algebra_1 + real_normed_algebra + | 
| 25062 | 784 | assumes norm_one [simp]: "norm 1 = 1" | 
| 62101 | 785 | |
| 63545 | 786 | lemma (in real_normed_algebra_1) scaleR_power [simp]: "(scaleR x y) ^ n = scaleR (x^n) (y^n)" | 
| 787 | by (induct n) (simp_all add: scaleR_one scaleR_scaleR mult_ac) | |
| 22852 | 788 | |
| 24588 | 789 | class real_normed_div_algebra = real_div_algebra + real_normed_vector + | 
| 25062 | 790 | assumes norm_mult: "norm (x * y) = norm x * norm y" | 
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 791 | |
| 78656 
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
 paulson <lp15@cam.ac.uk> parents: 
77221diff
changeset | 792 | lemma divideR_right: | 
| 
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
 paulson <lp15@cam.ac.uk> parents: 
77221diff
changeset | 793 | fixes x y :: "'a::real_normed_vector" | 
| 
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
 paulson <lp15@cam.ac.uk> parents: 
77221diff
changeset | 794 | shows "r \<noteq> 0 \<Longrightarrow> y = x /\<^sub>R r \<longleftrightarrow> r *\<^sub>R y = x" | 
| 
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
 paulson <lp15@cam.ac.uk> parents: 
77221diff
changeset | 795 | by auto | 
| 
4da1e18a9633
Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
 paulson <lp15@cam.ac.uk> parents: 
77221diff
changeset | 796 | |
| 24588 | 797 | class real_normed_field = real_field + real_normed_div_algebra | 
| 20584 
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
 huffman parents: 
20560diff
changeset | 798 | |
| 22852 | 799 | instance real_normed_div_algebra < real_normed_algebra_1 | 
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 800 | proof | 
| 63545 | 801 | show "norm (x * y) \<le> norm x * norm y" for x y :: 'a | 
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 802 | by (simp add: norm_mult) | 
| 22852 | 803 | next | 
| 804 | have "norm (1 * 1::'a) = norm (1::'a) * norm (1::'a)" | |
| 805 | by (rule norm_mult) | |
| 63545 | 806 | then show "norm (1::'a) = 1" by simp | 
| 20554 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 807 | qed | 
| 
c433e78d4203
define new constant of_real for class real_algebra_1;
 huffman parents: 
20551diff
changeset | 808 | |
| 69512 | 809 | context real_normed_vector begin | 
| 810 | ||
| 811 | lemma norm_zero [simp]: "norm (0::'a) = 0" | |
| 63545 | 812 | by simp | 
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 813 | |
| 63545 | 814 | lemma zero_less_norm_iff [simp]: "norm x > 0 \<longleftrightarrow> x \<noteq> 0" | 
| 815 | by (simp add: order_less_le) | |
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 816 | |
| 63545 | 817 | lemma norm_not_less_zero [simp]: "\<not> norm x < 0" | 
| 818 | by (simp add: linorder_not_less) | |
| 20828 | 819 | |
| 63545 | 820 | lemma norm_le_zero_iff [simp]: "norm x \<le> 0 \<longleftrightarrow> x = 0" | 
| 821 | by (simp add: order_le_less) | |
| 20828 | 822 | |
| 63545 | 823 | lemma norm_minus_cancel [simp]: "norm (- x) = norm x" | 
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 824 | proof - | 
| 69512 | 825 | have "- 1 *\<^sub>R x = - (1 *\<^sub>R x)" | 
| 826 | unfolding add_eq_0_iff2[symmetric] scaleR_add_left[symmetric] | |
| 827 | using norm_eq_zero | |
| 828 | by fastforce | |
| 829 | then have "norm (- x) = norm (scaleR (- 1) x)" | |
| 830 | by (simp only: scaleR_one) | |
| 20533 | 831 | also have "\<dots> = \<bar>- 1\<bar> * norm x" | 
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 832 | by (rule norm_scaleR) | 
| 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 833 | finally show ?thesis by simp | 
| 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 834 | qed | 
| 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 835 | |
| 63545 | 836 | lemma norm_minus_commute: "norm (a - b) = norm (b - a)" | 
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 837 | proof - | 
| 22898 | 838 | have "norm (- (b - a)) = norm (b - a)" | 
| 839 | by (rule norm_minus_cancel) | |
| 63545 | 840 | then show ?thesis by simp | 
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 841 | qed | 
| 63545 | 842 | |
| 843 | lemma dist_add_cancel [simp]: "dist (a + b) (a + c) = dist b c" | |
| 844 | by (simp add: dist_norm) | |
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63040diff
changeset | 845 | |
| 63545 | 846 | lemma dist_add_cancel2 [simp]: "dist (b + a) (c + a) = dist b c" | 
| 847 | by (simp add: dist_norm) | |
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63040diff
changeset | 848 | |
| 69512 | 849 | lemma norm_uminus_minus: "norm (- x - y) = norm (x + y)" | 
| 61524 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61169diff
changeset | 850 | by (subst (2) norm_minus_cancel[symmetric], subst minus_add_distrib) simp | 
| 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61169diff
changeset | 851 | |
| 63545 | 852 | lemma norm_triangle_ineq2: "norm a - norm b \<le> norm (a - b)" | 
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 853 | proof - | 
| 20533 | 854 | have "norm (a - b + b) \<le> norm (a - b) + norm b" | 
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 855 | by (rule norm_triangle_ineq) | 
| 63545 | 856 | then show ?thesis by simp | 
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 857 | qed | 
| 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 858 | |
| 63545 | 859 | lemma norm_triangle_ineq3: "\<bar>norm a - norm b\<bar> \<le> norm (a - b)" | 
| 68594 | 860 | proof - | 
| 861 | have "norm a - norm b \<le> norm (a - b)" | |
| 862 | by (simp add: norm_triangle_ineq2) | |
| 863 | moreover have "norm b - norm a \<le> norm (a - b)" | |
| 864 | by (metis norm_minus_commute norm_triangle_ineq2) | |
| 865 | ultimately show ?thesis | |
| 866 | by (simp add: abs_le_iff) | |
| 867 | qed | |
| 20584 
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
 huffman parents: 
20560diff
changeset | 868 | |
| 63545 | 869 | lemma norm_triangle_ineq4: "norm (a - b) \<le> norm a + norm b" | 
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 870 | proof - | 
| 22898 | 871 | have "norm (a + - b) \<le> norm a + norm (- b)" | 
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 872 | by (rule norm_triangle_ineq) | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53600diff
changeset | 873 | then show ?thesis by simp | 
| 22898 | 874 | qed | 
| 875 | ||
| 69512 | 876 | lemma norm_triangle_le_diff: "norm x + norm y \<le> e \<Longrightarrow> norm (x - y) \<le> e" | 
| 66422 | 877 | by (meson norm_triangle_ineq4 order_trans) | 
| 66420 | 878 | |
| 63545 | 879 | lemma norm_diff_ineq: "norm a - norm b \<le> norm (a + b)" | 
| 22898 | 880 | proof - | 
| 881 | have "norm a - norm (- b) \<le> norm (a - - b)" | |
| 882 | by (rule norm_triangle_ineq2) | |
| 63545 | 883 | then show ?thesis by simp | 
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 884 | qed | 
| 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 885 | |
| 69513 | 886 | lemma norm_triangle_sub: "norm x \<le> norm y + norm (x - y)" | 
| 887 | using norm_triangle_ineq[of "y" "x - y"] by (simp add: field_simps) | |
| 888 | ||
| 889 | lemma norm_triangle_le: "norm x + norm y \<le> e \<Longrightarrow> norm (x + y) \<le> e" | |
| 890 | by (rule norm_triangle_ineq [THEN order_trans]) | |
| 891 | ||
| 892 | lemma norm_triangle_lt: "norm x + norm y < e \<Longrightarrow> norm (x + y) < e" | |
| 893 | by (rule norm_triangle_ineq [THEN le_less_trans]) | |
| 894 | ||
| 63545 | 895 | lemma norm_add_leD: "norm (a + b) \<le> c \<Longrightarrow> norm b \<le> norm a + c" | 
| 69512 | 896 | by (metis ab_semigroup_add_class.add.commute add_commute diff_le_eq norm_diff_ineq order_trans) | 
| 61762 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 paulson <lp15@cam.ac.uk> parents: 
61649diff
changeset | 897 | |
| 63545 | 898 | lemma norm_diff_triangle_ineq: "norm ((a + b) - (c + d)) \<le> norm (a - c) + norm (b - d)" | 
| 20551 | 899 | proof - | 
| 900 | have "norm ((a + b) - (c + d)) = norm ((a - c) + (b - d))" | |
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53600diff
changeset | 901 | by (simp add: algebra_simps) | 
| 20551 | 902 | also have "\<dots> \<le> norm (a - c) + norm (b - d)" | 
| 903 | by (rule norm_triangle_ineq) | |
| 904 | finally show ?thesis . | |
| 905 | qed | |
| 906 | ||
| 69512 | 907 | lemma norm_diff_triangle_le: "norm (x - z) \<le> e1 + e2" | 
| 908 | if "norm (x - y) \<le> e1" "norm (y - z) \<le> e2" | |
| 909 | proof - | |
| 910 | have "norm (x - (y + z - y)) \<le> norm (x - y) + norm (y - z)" | |
| 911 | using norm_diff_triangle_ineq that diff_diff_eq2 by presburger | |
| 912 | with that show ?thesis by simp | |
| 913 | qed | |
| 60800 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 914 | |
| 69512 | 915 | lemma norm_diff_triangle_less: "norm (x - z) < e1 + e2" | 
| 916 | if "norm (x - y) < e1" "norm (y - z) < e2" | |
| 917 | proof - | |
| 918 | have "norm (x - z) \<le> norm (x - y) + norm (y - z)" | |
| 919 | by (metis norm_diff_triangle_ineq add_diff_cancel_left' diff_diff_eq2) | |
| 920 | with that show ?thesis by auto | |
| 921 | qed | |
| 60800 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 922 | |
| 60026 
41d81b4a0a21
Restored LIMSEQ_def as legacy binding. [The other changes are whitespace only.]
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 923 | lemma norm_triangle_mono: | 
| 69512 | 924 | "norm a \<le> r \<Longrightarrow> norm b \<le> s \<Longrightarrow> norm (a + b) \<le> r + s" | 
| 925 | by (metis (mono_tags) add_mono_thms_linordered_semiring(1) norm_triangle_ineq order.trans) | |
| 55719 
cdddd073bff8
Lemmas about Reals, norm, etc., and cleaner variants of existing ones
 paulson <lp15@cam.ac.uk> parents: 
54890diff
changeset | 926 | |
| 69512 | 927 | lemma norm_sum: "norm (sum f A) \<le> (\<Sum>i\<in>A. norm (f i))" | 
| 928 | for f::"'b \<Rightarrow> 'a" | |
| 56194 | 929 | by (induct A rule: infinite_finite_induct) (auto intro: norm_triangle_mono) | 
| 930 | ||
| 69512 | 931 | lemma sum_norm_le: "norm (sum f S) \<le> sum g S" | 
| 932 | if "\<And>x. x \<in> S \<Longrightarrow> norm (f x) \<le> g x" | |
| 933 | for f::"'b \<Rightarrow> 'a" | |
| 934 | by (rule order_trans [OF norm_sum sum_mono]) (simp add: that) | |
| 56369 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56194diff
changeset | 935 | |
| 63545 | 936 | lemma abs_norm_cancel [simp]: "\<bar>norm a\<bar> = norm a" | 
| 937 | by (rule abs_of_nonneg [OF norm_ge_zero]) | |
| 22857 | 938 | |
| 69513 | 939 | lemma sum_norm_bound: | 
| 940 | "norm (sum f S) \<le> of_nat (card S)*K" | |
| 941 | if "\<And>x. x \<in> S \<Longrightarrow> norm (f x) \<le> K" | |
| 942 | for f :: "'b \<Rightarrow> 'a" | |
| 943 | using sum_norm_le[OF that] sum_constant[symmetric] | |
| 944 | by simp | |
| 945 | ||
| 63545 | 946 | lemma norm_add_less: "norm x < r \<Longrightarrow> norm y < s \<Longrightarrow> norm (x + y) < r + s" | 
| 947 | by (rule order_le_less_trans [OF norm_triangle_ineq add_strict_mono]) | |
| 22880 | 948 | |
| 69512 | 949 | end | 
| 950 | ||
| 951 | lemma dist_scaleR [simp]: "dist (x *\<^sub>R a) (y *\<^sub>R a) = \<bar>x - y\<bar> * norm a" | |
| 952 | for a :: "'a::real_normed_vector" | |
| 953 | by (metis dist_norm norm_scaleR scaleR_left.diff) | |
| 954 | ||
| 63545 | 955 | lemma norm_mult_less: "norm x < r \<Longrightarrow> norm y < s \<Longrightarrow> norm (x * y) < r * s" | 
| 956 | for x y :: "'a::real_normed_algebra" | |
| 957 | by (rule order_le_less_trans [OF norm_mult_ineq]) (simp add: mult_strict_mono') | |
| 22880 | 958 | |
| 63545 | 959 | lemma norm_of_real [simp]: "norm (of_real r :: 'a::real_normed_algebra_1) = \<bar>r\<bar>" | 
| 960 | by (simp add: of_real_def) | |
| 20560 | 961 | |
| 63545 | 962 | lemma norm_numeral [simp]: "norm (numeral w::'a::real_normed_algebra_1) = numeral w" | 
| 963 | by (subst of_real_numeral [symmetric], subst norm_of_real, simp) | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46868diff
changeset | 964 | |
| 63545 | 965 | lemma norm_neg_numeral [simp]: "norm (- numeral w::'a::real_normed_algebra_1) = numeral w" | 
| 966 | by (subst of_real_neg_numeral [symmetric], subst norm_of_real, simp) | |
| 22876 
2b4c831ceca7
add lemmas norm_number_of, norm_of_int, norm_of_nat
 huffman parents: 
22857diff
changeset | 967 | |
| 63545 | 968 | lemma norm_of_real_add1 [simp]: "norm (of_real x + 1 :: 'a :: real_normed_div_algebra) = \<bar>x + 1\<bar>" | 
| 62379 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62368diff
changeset | 969 | by (metis norm_of_real of_real_1 of_real_add) | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62368diff
changeset | 970 | |
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62368diff
changeset | 971 | lemma norm_of_real_addn [simp]: | 
| 63545 | 972 | "norm (of_real x + numeral b :: 'a :: real_normed_div_algebra) = \<bar>x + numeral b\<bar>" | 
| 62379 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62368diff
changeset | 973 | by (metis norm_of_real of_real_add of_real_numeral) | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62368diff
changeset | 974 | |
| 63545 | 975 | lemma norm_of_int [simp]: "norm (of_int z::'a::real_normed_algebra_1) = \<bar>of_int z\<bar>" | 
| 976 | by (subst of_real_of_int_eq [symmetric], rule norm_of_real) | |
| 22876 
2b4c831ceca7
add lemmas norm_number_of, norm_of_int, norm_of_nat
 huffman parents: 
22857diff
changeset | 977 | |
| 63545 | 978 | lemma norm_of_nat [simp]: "norm (of_nat n::'a::real_normed_algebra_1) = of_nat n" | 
| 68594 | 979 | by (metis abs_of_nat norm_of_real of_real_of_nat_eq) | 
| 22876 
2b4c831ceca7
add lemmas norm_number_of, norm_of_int, norm_of_nat
 huffman parents: 
22857diff
changeset | 980 | |
| 63545 | 981 | lemma nonzero_norm_inverse: "a \<noteq> 0 \<Longrightarrow> norm (inverse a) = inverse (norm a)" | 
| 982 | for a :: "'a::real_normed_div_algebra" | |
| 68594 | 983 | by (metis inverse_unique norm_mult norm_one right_inverse) | 
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 984 | |
| 63545 | 985 | lemma norm_inverse: "norm (inverse a) = inverse (norm a)" | 
| 986 |   for a :: "'a::{real_normed_div_algebra,division_ring}"
 | |
| 68594 | 987 | by (metis inverse_zero nonzero_norm_inverse norm_zero) | 
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 988 | |
| 63545 | 989 | lemma nonzero_norm_divide: "b \<noteq> 0 \<Longrightarrow> norm (a / b) = norm a / norm b" | 
| 990 | for a b :: "'a::real_normed_field" | |
| 991 | by (simp add: divide_inverse norm_mult nonzero_norm_inverse) | |
| 20584 
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
 huffman parents: 
20560diff
changeset | 992 | |
| 63545 | 993 | lemma norm_divide: "norm (a / b) = norm a / norm b" | 
| 994 |   for a b :: "'a::{real_normed_field,field}"
 | |
| 995 | by (simp add: divide_inverse norm_mult norm_inverse) | |
| 20584 
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
 huffman parents: 
20560diff
changeset | 996 | |
| 77221 
0cdb384bf56a
More new theorems from the number theory development
 paulson <lp15@cam.ac.uk> parents: 
77138diff
changeset | 997 | lemma dist_divide_right: "dist (a/c) (b/c) = dist a b / norm c" for c :: "'a :: real_normed_field" | 
| 
0cdb384bf56a
More new theorems from the number theory development
 paulson <lp15@cam.ac.uk> parents: 
77138diff
changeset | 998 | by (metis diff_divide_distrib dist_norm norm_divide) | 
| 
0cdb384bf56a
More new theorems from the number theory development
 paulson <lp15@cam.ac.uk> parents: 
77138diff
changeset | 999 | |
| 68615 | 1000 | lemma norm_inverse_le_norm: | 
| 1001 | fixes x :: "'a::real_normed_div_algebra" | |
| 1002 | shows "r \<le> norm x \<Longrightarrow> 0 < r \<Longrightarrow> norm (inverse x) \<le> inverse r" | |
| 1003 | by (simp add: le_imp_inverse_le norm_inverse) | |
| 1004 | ||
| 63545 | 1005 | lemma norm_power_ineq: "norm (x ^ n) \<le> norm x ^ n" | 
| 1006 | for x :: "'a::real_normed_algebra_1" | |
| 22852 | 1007 | proof (induct n) | 
| 63545 | 1008 | case 0 | 
| 1009 | show "norm (x ^ 0) \<le> norm x ^ 0" by simp | |
| 22852 | 1010 | next | 
| 1011 | case (Suc n) | |
| 1012 | have "norm (x * x ^ n) \<le> norm x * norm (x ^ n)" | |
| 1013 | by (rule norm_mult_ineq) | |
| 1014 | also from Suc have "\<dots> \<le> norm x * norm x ^ n" | |
| 1015 | using norm_ge_zero by (rule mult_left_mono) | |
| 1016 | finally show "norm (x ^ Suc n) \<le> norm x ^ Suc n" | |
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 1017 | by simp | 
| 22852 | 1018 | qed | 
| 1019 | ||
| 63545 | 1020 | lemma norm_power: "norm (x ^ n) = norm x ^ n" | 
| 1021 | for x :: "'a::real_normed_div_algebra" | |
| 1022 | by (induct n) (simp_all add: norm_mult) | |
| 20684 | 1023 | |
| 71837 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 1024 | lemma norm_power_int: "norm (power_int x n) = power_int (norm x) n" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 1025 | for x :: "'a::real_normed_div_algebra" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 1026 | by (cases n rule: int_cases4) (auto simp: norm_power power_int_minus norm_inverse) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 1027 | |
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62623diff
changeset | 1028 | lemma power_eq_imp_eq_norm: | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62623diff
changeset | 1029 | fixes w :: "'a::real_normed_div_algebra" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62623diff
changeset | 1030 | assumes eq: "w ^ n = z ^ n" and "n > 0" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62623diff
changeset | 1031 | shows "norm w = norm z" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62623diff
changeset | 1032 | proof - | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62623diff
changeset | 1033 | have "norm w ^ n = norm z ^ n" | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62623diff
changeset | 1034 | by (metis (no_types) eq norm_power) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62623diff
changeset | 1035 | then show ?thesis | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62623diff
changeset | 1036 | using assms by (force intro: power_eq_imp_eq_base) | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62623diff
changeset | 1037 | qed | 
| 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62623diff
changeset | 1038 | |
| 68465 
e699ca8e22b7
New material in support of quaternions
 paulson <lp15@cam.ac.uk> parents: 
68397diff
changeset | 1039 | lemma power_eq_1_iff: | 
| 
e699ca8e22b7
New material in support of quaternions
 paulson <lp15@cam.ac.uk> parents: 
68397diff
changeset | 1040 | fixes w :: "'a::real_normed_div_algebra" | 
| 
e699ca8e22b7
New material in support of quaternions
 paulson <lp15@cam.ac.uk> parents: 
68397diff
changeset | 1041 | shows "w ^ n = 1 \<Longrightarrow> norm w = 1 \<or> n = 0" | 
| 
e699ca8e22b7
New material in support of quaternions
 paulson <lp15@cam.ac.uk> parents: 
68397diff
changeset | 1042 | by (metis norm_one power_0_left power_eq_0_iff power_eq_imp_eq_norm power_one) | 
| 
e699ca8e22b7
New material in support of quaternions
 paulson <lp15@cam.ac.uk> parents: 
68397diff
changeset | 1043 | |
| 63545 | 1044 | lemma norm_mult_numeral1 [simp]: "norm (numeral w * a) = numeral w * norm a" | 
| 1045 |   for a b :: "'a::{real_normed_field,field}"
 | |
| 1046 | by (simp add: norm_mult) | |
| 60762 | 1047 | |
| 63545 | 1048 | lemma norm_mult_numeral2 [simp]: "norm (a * numeral w) = norm a * numeral w" | 
| 1049 |   for a b :: "'a::{real_normed_field,field}"
 | |
| 1050 | by (simp add: norm_mult) | |
| 60762 | 1051 | |
| 63545 | 1052 | lemma norm_divide_numeral [simp]: "norm (a / numeral w) = norm a / numeral w" | 
| 1053 |   for a b :: "'a::{real_normed_field,field}"
 | |
| 1054 | by (simp add: norm_divide) | |
| 60762 | 1055 | |
| 1056 | lemma norm_of_real_diff [simp]: | |
| 63545 | 1057 | "norm (of_real b - of_real a :: 'a::real_normed_algebra_1) \<le> \<bar>b - a\<bar>" | 
| 60762 | 1058 | by (metis norm_of_real of_real_diff order_refl) | 
| 1059 | ||
| 63545 | 1060 | text \<open>Despite a superficial resemblance, \<open>norm_eq_1\<close> is not relevant.\<close> | 
| 59613 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
59587diff
changeset | 1061 | lemma square_norm_one: | 
| 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
59587diff
changeset | 1062 | fixes x :: "'a::real_normed_div_algebra" | 
| 63545 | 1063 | assumes "x\<^sup>2 = 1" | 
| 1064 | shows "norm x = 1" | |
| 59613 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
59587diff
changeset | 1065 | by (metis assms norm_minus_cancel norm_one power2_eq_1_iff) | 
| 
7103019278f0
The function frac. Various lemmas about limits, series, the exp function, etc.
 paulson <lp15@cam.ac.uk> parents: 
59587diff
changeset | 1066 | |
| 63545 | 1067 | lemma norm_less_p1: "norm x < norm (of_real (norm x) + 1 :: 'a)" | 
| 1068 | for x :: "'a::real_normed_algebra_1" | |
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
59613diff
changeset | 1069 | proof - | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
59613diff
changeset | 1070 | have "norm x < norm (of_real (norm x + 1) :: 'a)" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
59613diff
changeset | 1071 | by (simp add: of_real_def) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
59613diff
changeset | 1072 | then show ?thesis | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
59613diff
changeset | 1073 | by simp | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
59613diff
changeset | 1074 | qed | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
59613diff
changeset | 1075 | |
| 64272 | 1076 | lemma prod_norm: "prod (\<lambda>x. norm (f x)) A = norm (prod f A)" | 
| 63545 | 1077 |   for f :: "'a \<Rightarrow> 'b::{comm_semiring_1,real_normed_div_algebra}"
 | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1078 | by (induct A rule: infinite_finite_induct) (auto simp: norm_mult) | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1079 | |
| 64272 | 1080 | lemma norm_prod_le: | 
| 1081 |   "norm (prod f A) \<le> (\<Prod>a\<in>A. norm (f a :: 'a :: {real_normed_algebra_1,comm_monoid_mult}))"
 | |
| 63545 | 1082 | proof (induct A rule: infinite_finite_induct) | 
| 1083 | case empty | |
| 1084 | then show ?case by simp | |
| 1085 | next | |
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1086 | case (insert a A) | 
| 64272 | 1087 | then have "norm (prod f (insert a A)) \<le> norm (f a) * norm (prod f A)" | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1088 | by (simp add: norm_mult_ineq) | 
| 64272 | 1089 | also have "norm (prod f A) \<le> (\<Prod>a\<in>A. norm (f a))" | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1090 | by (rule insert) | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1091 | finally show ?case | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1092 | by (simp add: insert mult_left_mono) | 
| 63545 | 1093 | next | 
| 1094 | case infinite | |
| 1095 | then show ?case by simp | |
| 1096 | qed | |
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1097 | |
| 64272 | 1098 | lemma norm_prod_diff: | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1099 |   fixes z w :: "'i \<Rightarrow> 'a::{real_normed_algebra_1, comm_monoid_mult}"
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1100 | shows "(\<And>i. i \<in> I \<Longrightarrow> norm (z i) \<le> 1) \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> norm (w i) \<le> 1) \<Longrightarrow> | 
| 60026 
41d81b4a0a21
Restored LIMSEQ_def as legacy binding. [The other changes are whitespace only.]
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 1101 | norm ((\<Prod>i\<in>I. z i) - (\<Prod>i\<in>I. w i)) \<le> (\<Sum>i\<in>I. norm (z i - w i))" | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1102 | proof (induction I rule: infinite_finite_induct) | 
| 63545 | 1103 | case empty | 
| 1104 | then show ?case by simp | |
| 1105 | next | |
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1106 | case (insert i I) | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1107 | note insert.hyps[simp] | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1108 | |
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1109 | have "norm ((\<Prod>i\<in>insert i I. z i) - (\<Prod>i\<in>insert i I. w i)) = | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1110 | norm ((\<Prod>i\<in>I. z i) * (z i - w i) + ((\<Prod>i\<in>I. z i) - (\<Prod>i\<in>I. w i)) * w i)" | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1111 | (is "_ = norm (?t1 + ?t2)") | 
| 68594 | 1112 | by (auto simp: field_simps) | 
| 63545 | 1113 | also have "\<dots> \<le> norm ?t1 + norm ?t2" | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1114 | by (rule norm_triangle_ineq) | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1115 | also have "norm ?t1 \<le> norm (\<Prod>i\<in>I. z i) * norm (z i - w i)" | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1116 | by (rule norm_mult_ineq) | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1117 | also have "\<dots> \<le> (\<Prod>i\<in>I. norm (z i)) * norm(z i - w i)" | 
| 64272 | 1118 | by (rule mult_right_mono) (auto intro: norm_prod_le) | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1119 | also have "(\<Prod>i\<in>I. norm (z i)) \<le> (\<Prod>i\<in>I. 1)" | 
| 64272 | 1120 | by (intro prod_mono) (auto intro!: insert) | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1121 | also have "norm ?t2 \<le> norm ((\<Prod>i\<in>I. z i) - (\<Prod>i\<in>I. w i)) * norm (w i)" | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1122 | by (rule norm_mult_ineq) | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1123 | also have "norm (w i) \<le> 1" | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1124 | by (auto intro: insert) | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1125 | also have "norm ((\<Prod>i\<in>I. z i) - (\<Prod>i\<in>I. w i)) \<le> (\<Sum>i\<in>I. norm (z i - w i))" | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1126 | using insert by auto | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1127 | finally show ?case | 
| 68594 | 1128 | by (auto simp: ac_simps mult_right_mono mult_left_mono) | 
| 63545 | 1129 | next | 
| 1130 | case infinite | |
| 1131 | then show ?case by simp | |
| 1132 | qed | |
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1133 | |
| 60026 
41d81b4a0a21
Restored LIMSEQ_def as legacy binding. [The other changes are whitespace only.]
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 1134 | lemma norm_power_diff: | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1135 |   fixes z w :: "'a::{real_normed_algebra_1, comm_monoid_mult}"
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1136 | assumes "norm z \<le> 1" "norm w \<le> 1" | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1137 | shows "norm (z^m - w^m) \<le> m * norm (z - w)" | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1138 | proof - | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1139 | have "norm (z^m - w^m) = norm ((\<Prod> i < m. z) - (\<Prod> i < m. w))" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 1140 | by simp | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1141 | also have "\<dots> \<le> (\<Sum>i<m. norm (z - w))" | 
| 68594 | 1142 | by (intro norm_prod_diff) (auto simp: assms) | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1143 | also have "\<dots> = m * norm (z - w)" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 1144 | by simp | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 1145 | finally show ?thesis . | 
| 55719 
cdddd073bff8
Lemmas about Reals, norm, etc., and cleaner variants of existing ones
 paulson <lp15@cam.ac.uk> parents: 
54890diff
changeset | 1146 | qed | 
| 
cdddd073bff8
Lemmas about Reals, norm, etc., and cleaner variants of existing ones
 paulson <lp15@cam.ac.uk> parents: 
54890diff
changeset | 1147 | |
| 60758 | 1148 | subsection \<open>Metric spaces\<close> | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1149 | |
| 62101 | 1150 | class metric_space = uniformity_dist + open_uniformity + | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1151 | assumes dist_eq_0_iff [simp]: "dist x y = 0 \<longleftrightarrow> x = y" | 
| 63545 | 1152 | and dist_triangle2: "dist x y \<le> dist x z + dist y z" | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1153 | begin | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1154 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1155 | lemma dist_self [simp]: "dist x x = 0" | 
| 63545 | 1156 | by simp | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1157 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1158 | lemma zero_le_dist [simp]: "0 \<le> dist x y" | 
| 63545 | 1159 | using dist_triangle2 [of x x y] by simp | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1160 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1161 | lemma zero_less_dist_iff: "0 < dist x y \<longleftrightarrow> x \<noteq> y" | 
| 63545 | 1162 | by (simp add: less_le) | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1163 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1164 | lemma dist_not_less_zero [simp]: "\<not> dist x y < 0" | 
| 63545 | 1165 | by (simp add: not_less) | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1166 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1167 | lemma dist_le_zero_iff [simp]: "dist x y \<le> 0 \<longleftrightarrow> x = y" | 
| 63545 | 1168 | by (simp add: le_less) | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1169 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1170 | lemma dist_commute: "dist x y = dist y x" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1171 | proof (rule order_antisym) | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1172 | show "dist x y \<le> dist y x" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1173 | using dist_triangle2 [of x y x] by simp | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1174 | show "dist y x \<le> dist x y" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1175 | using dist_triangle2 [of y x y] by simp | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1176 | qed | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1177 | |
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 1178 | lemma dist_commute_lessI: "dist y x < e \<Longrightarrow> dist x y < e" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 1179 | by (simp add: dist_commute) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 1180 | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1181 | lemma dist_triangle: "dist x z \<le> dist x y + dist y z" | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 1182 | using dist_triangle2 [of x z y] by (simp add: dist_commute) | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1183 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1184 | lemma dist_triangle3: "dist x y \<le> dist a x + dist a y" | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 1185 | using dist_triangle2 [of x y a] by (simp add: dist_commute) | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1186 | |
| 68721 | 1187 | lemma abs_dist_diff_le: "\<bar>dist a b - dist b c\<bar> \<le> dist a c" | 
| 1188 | using dist_triangle3[of b c a] dist_triangle2[of a b c] by simp | |
| 1189 | ||
| 63545 | 1190 | lemma dist_pos_lt: "x \<noteq> y \<Longrightarrow> 0 < dist x y" | 
| 1191 | by (simp add: zero_less_dist_iff) | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1192 | |
| 63545 | 1193 | lemma dist_nz: "x \<noteq> y \<longleftrightarrow> 0 < dist x y" | 
| 1194 | by (simp add: zero_less_dist_iff) | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1195 | |
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
62049diff
changeset | 1196 | declare dist_nz [symmetric, simp] | 
| 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
62049diff
changeset | 1197 | |
| 63545 | 1198 | lemma dist_triangle_le: "dist x z + dist y z \<le> e \<Longrightarrow> dist x y \<le> e" | 
| 1199 | by (rule order_trans [OF dist_triangle2]) | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1200 | |
| 63545 | 1201 | lemma dist_triangle_lt: "dist x z + dist y z < e \<Longrightarrow> dist x y < e" | 
| 1202 | by (rule le_less_trans [OF dist_triangle2]) | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1203 | |
| 63545 | 1204 | lemma dist_triangle_less_add: "dist x1 y < e1 \<Longrightarrow> dist x2 y < e2 \<Longrightarrow> dist x1 x2 < e1 + e2" | 
| 1205 | by (rule dist_triangle_lt [where z=y]) simp | |
| 62948 
7700f467892b
lots of new theorems for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
62623diff
changeset | 1206 | |
| 63545 | 1207 | lemma dist_triangle_half_l: "dist x1 y < e / 2 \<Longrightarrow> dist x2 y < e / 2 \<Longrightarrow> dist x1 x2 < e" | 
| 1208 | by (rule dist_triangle_lt [where z=y]) simp | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1209 | |
| 63545 | 1210 | lemma dist_triangle_half_r: "dist y x1 < e / 2 \<Longrightarrow> dist y x2 < e / 2 \<Longrightarrow> dist x1 x2 < e" | 
| 1211 | by (rule dist_triangle_half_l) (simp_all add: dist_commute) | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1212 | |
| 65036 
ab7e11730ad8
Some new lemmas. Existing lemmas modified to use uniform_limit rather than its expansion
 paulson <lp15@cam.ac.uk> parents: 
64788diff
changeset | 1213 | lemma dist_triangle_third: | 
| 
ab7e11730ad8
Some new lemmas. Existing lemmas modified to use uniform_limit rather than its expansion
 paulson <lp15@cam.ac.uk> parents: 
64788diff
changeset | 1214 | assumes "dist x1 x2 < e/3" "dist x2 x3 < e/3" "dist x3 x4 < e/3" | 
| 
ab7e11730ad8
Some new lemmas. Existing lemmas modified to use uniform_limit rather than its expansion
 paulson <lp15@cam.ac.uk> parents: 
64788diff
changeset | 1215 | shows "dist x1 x4 < e" | 
| 
ab7e11730ad8
Some new lemmas. Existing lemmas modified to use uniform_limit rather than its expansion
 paulson <lp15@cam.ac.uk> parents: 
64788diff
changeset | 1216 | proof - | 
| 
ab7e11730ad8
Some new lemmas. Existing lemmas modified to use uniform_limit rather than its expansion
 paulson <lp15@cam.ac.uk> parents: 
64788diff
changeset | 1217 | have "dist x1 x3 < e/3 + e/3" | 
| 
ab7e11730ad8
Some new lemmas. Existing lemmas modified to use uniform_limit rather than its expansion
 paulson <lp15@cam.ac.uk> parents: 
64788diff
changeset | 1218 | by (metis assms(1) assms(2) dist_commute dist_triangle_less_add) | 
| 
ab7e11730ad8
Some new lemmas. Existing lemmas modified to use uniform_limit rather than its expansion
 paulson <lp15@cam.ac.uk> parents: 
64788diff
changeset | 1219 | then have "dist x1 x4 < (e/3 + e/3) + e/3" | 
| 
ab7e11730ad8
Some new lemmas. Existing lemmas modified to use uniform_limit rather than its expansion
 paulson <lp15@cam.ac.uk> parents: 
64788diff
changeset | 1220 | by (metis assms(3) dist_commute dist_triangle_less_add) | 
| 
ab7e11730ad8
Some new lemmas. Existing lemmas modified to use uniform_limit rather than its expansion
 paulson <lp15@cam.ac.uk> parents: 
64788diff
changeset | 1221 | then show ?thesis | 
| 
ab7e11730ad8
Some new lemmas. Existing lemmas modified to use uniform_limit rather than its expansion
 paulson <lp15@cam.ac.uk> parents: 
64788diff
changeset | 1222 | by simp | 
| 
ab7e11730ad8
Some new lemmas. Existing lemmas modified to use uniform_limit rather than its expansion
 paulson <lp15@cam.ac.uk> parents: 
64788diff
changeset | 1223 | qed | 
| 68532 
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
 paulson <lp15@cam.ac.uk> parents: 
68499diff
changeset | 1224 | |
| 62101 | 1225 | subclass uniform_space | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1226 | proof | 
| 63545 | 1227 | fix E x | 
| 1228 | assume "eventually E uniformity" | |
| 62101 | 1229 | then obtain e where E: "0 < e" "\<And>x y. dist x y < e \<Longrightarrow> E (x, y)" | 
| 63545 | 1230 | by (auto simp: eventually_uniformity_metric) | 
| 62101 | 1231 | then show "E (x, x)" "\<forall>\<^sub>F (x, y) in uniformity. E (y, x)" | 
| 63545 | 1232 | by (auto simp: eventually_uniformity_metric dist_commute) | 
| 62101 | 1233 | show "\<exists>D. eventually D uniformity \<and> (\<forall>x y z. D (x, y) \<longrightarrow> D (y, z) \<longrightarrow> E (x, z))" | 
| 63545 | 1234 | using E dist_triangle_half_l[where e=e] | 
| 1235 | unfolding eventually_uniformity_metric | |
| 62101 | 1236 | by (intro exI[of _ "\<lambda>(x, y). dist x y < e / 2"] exI[of _ "e/2"] conjI) | 
| 63545 | 1237 | (auto simp: dist_commute) | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1238 | qed | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1239 | |
| 62101 | 1240 | lemma open_dist: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)" | 
| 63545 | 1241 | by (simp add: dist_commute open_uniformity eventually_uniformity_metric) | 
| 62101 | 1242 | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1243 | lemma open_ball: "open {y. dist x y < d}"
 | 
| 63545 | 1244 | unfolding open_dist | 
| 1245 | proof (intro ballI) | |
| 1246 | fix y | |
| 1247 |   assume *: "y \<in> {y. dist x y < d}"
 | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1248 |   then show "\<exists>e>0. \<forall>z. dist z y < e \<longrightarrow> z \<in> {y. dist x y < d}"
 | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1249 | by (auto intro!: exI[of _ "d - dist x y"] simp: field_simps dist_triangle_lt) | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1250 | qed | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1251 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1252 | subclass first_countable_topology | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1253 | proof | 
| 60026 
41d81b4a0a21
Restored LIMSEQ_def as legacy binding. [The other changes are whitespace only.]
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 1254 | fix x | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1255 | show "\<exists>A::nat \<Rightarrow> 'a set. (\<forall>i. x \<in> A i \<and> open (A i)) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>i. A i \<subseteq> S))" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1256 |   proof (safe intro!: exI[of _ "\<lambda>n. {y. dist x y < inverse (Suc n)}"])
 | 
| 63545 | 1257 | fix S | 
| 1258 | assume "open S" "x \<in> S" | |
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
52381diff
changeset | 1259 |     then obtain e where e: "0 < e" and "{y. dist x y < e} \<subseteq> S"
 | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1260 | by (auto simp: open_dist subset_eq dist_commute) | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1261 | moreover | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
52381diff
changeset | 1262 | from e obtain i where "inverse (Suc i) < e" | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1263 | by (auto dest!: reals_Archimedean) | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1264 |     then have "{y. dist x y < inverse (Suc i)} \<subseteq> {y. dist x y < e}"
 | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1265 | by auto | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1266 |     ultimately show "\<exists>i. {y. dist x y < inverse (Suc i)} \<subseteq> S"
 | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1267 | by blast | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1268 | qed (auto intro: open_ball) | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1269 | qed | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1270 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1271 | end | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1272 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1273 | instance metric_space \<subseteq> t2_space | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1274 | proof | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1275 | fix x y :: "'a::metric_space" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1276 | assume xy: "x \<noteq> y" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1277 |   let ?U = "{y'. dist x y' < dist x y / 2}"
 | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1278 |   let ?V = "{x'. dist y x' < dist x y / 2}"
 | 
| 63545 | 1279 | have *: "d x z \<le> d x y + d y z \<Longrightarrow> d y z = d z y \<Longrightarrow> \<not> (d x y * 2 < d x z \<and> d z y * 2 < d x z)" | 
| 1280 | for d :: "'a \<Rightarrow> 'a \<Rightarrow> real" and x y z :: 'a | |
| 1281 | by arith | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1282 |   have "open ?U \<and> open ?V \<and> x \<in> ?U \<and> y \<in> ?V \<and> ?U \<inter> ?V = {}"
 | 
| 63545 | 1283 | using dist_pos_lt[OF xy] *[of dist, OF dist_triangle dist_commute] | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1284 | using open_ball[of _ "dist x y / 2"] by auto | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1285 |   then show "\<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
 | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1286 | by blast | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1287 | qed | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1288 | |
| 60758 | 1289 | text \<open>Every normed vector space is a metric space.\<close> | 
| 31289 | 1290 | instance real_normed_vector < metric_space | 
| 1291 | proof | |
| 63545 | 1292 | fix x y z :: 'a | 
| 1293 | show "dist x y = 0 \<longleftrightarrow> x = y" | |
| 1294 | by (simp add: dist_norm) | |
| 1295 | show "dist x y \<le> dist x z + dist y z" | |
| 1296 | using norm_triangle_ineq4 [of "x - z" "y - z"] by (simp add: dist_norm) | |
| 31289 | 1297 | qed | 
| 31285 
0a3f9ee4117c
generalize dist function to class real_normed_vector
 huffman parents: 
31017diff
changeset | 1298 | |
| 63545 | 1299 | |
| 60758 | 1300 | subsection \<open>Class instances for real numbers\<close> | 
| 31564 
d2abf6f6f619
subsection for real instances; new lemmas for open sets of reals
 huffman parents: 
31494diff
changeset | 1301 | |
| 
d2abf6f6f619
subsection for real instances; new lemmas for open sets of reals
 huffman parents: 
31494diff
changeset | 1302 | instantiation real :: real_normed_field | 
| 
d2abf6f6f619
subsection for real instances; new lemmas for open sets of reals
 huffman parents: 
31494diff
changeset | 1303 | begin | 
| 
d2abf6f6f619
subsection for real instances; new lemmas for open sets of reals
 huffman parents: 
31494diff
changeset | 1304 | |
| 63545 | 1305 | definition dist_real_def: "dist x y = \<bar>x - y\<bar>" | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1306 | |
| 62101 | 1307 | definition uniformity_real_def [code del]: | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69064diff
changeset | 1308 |   "(uniformity :: (real \<times> real) filter) = (INF e\<in>{0 <..}. principal {(x, y). dist x y < e})"
 | 
| 62101 | 1309 | |
| 52381 
63eec9cea2c7
pragmatic executability for instance real :: open
 haftmann parents: 
51775diff
changeset | 1310 | definition open_real_def [code del]: | 
| 62101 | 1311 | "open (U :: real set) \<longleftrightarrow> (\<forall>x\<in>U. eventually (\<lambda>(x', y). x' = x \<longrightarrow> y \<in> U) uniformity)" | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1312 | |
| 63545 | 1313 | definition real_norm_def [simp]: "norm r = \<bar>r\<bar>" | 
| 31564 
d2abf6f6f619
subsection for real instances; new lemmas for open sets of reals
 huffman parents: 
31494diff
changeset | 1314 | |
| 
d2abf6f6f619
subsection for real instances; new lemmas for open sets of reals
 huffman parents: 
31494diff
changeset | 1315 | instance | 
| 68594 | 1316 | by intro_classes (auto simp: abs_mult open_real_def dist_real_def sgn_real_def uniformity_real_def) | 
| 31564 
d2abf6f6f619
subsection for real instances; new lemmas for open sets of reals
 huffman parents: 
31494diff
changeset | 1317 | |
| 
d2abf6f6f619
subsection for real instances; new lemmas for open sets of reals
 huffman parents: 
31494diff
changeset | 1318 | end | 
| 
d2abf6f6f619
subsection for real instances; new lemmas for open sets of reals
 huffman parents: 
31494diff
changeset | 1319 | |
| 62102 
877463945ce9
fix code generation for uniformity: uniformity is a non-computable pure data.
 hoelzl parents: 
62101diff
changeset | 1320 | declare uniformity_Abort[where 'a=real, code] | 
| 
877463945ce9
fix code generation for uniformity: uniformity is a non-computable pure data.
 hoelzl parents: 
62101diff
changeset | 1321 | |
| 63545 | 1322 | lemma dist_of_real [simp]: "dist (of_real x :: 'a) (of_real y) = dist x y" | 
| 1323 | for a :: "'a::real_normed_div_algebra" | |
| 1324 | by (metis dist_norm norm_of_real of_real_diff real_norm_def) | |
| 60800 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 1325 | |
| 54890 
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
 haftmann parents: 
54863diff
changeset | 1326 | declare [[code abort: "open :: real set \<Rightarrow> bool"]] | 
| 52381 
63eec9cea2c7
pragmatic executability for instance real :: open
 haftmann parents: 
51775diff
changeset | 1327 | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1328 | instance real :: linorder_topology | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1329 | proof | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1330 | show "(open :: real set \<Rightarrow> bool) = generate_topology (range lessThan \<union> range greaterThan)" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1331 | proof (rule ext, safe) | 
| 63545 | 1332 | fix S :: "real set" | 
| 1333 | assume "open S" | |
| 53381 | 1334 | then obtain f where "\<forall>x\<in>S. 0 < f x \<and> (\<forall>y. dist y x < f x \<longrightarrow> y \<in> S)" | 
| 62101 | 1335 | unfolding open_dist bchoice_iff .. | 
| 71720 | 1336 |     then have *: "(\<Union>x\<in>S. {x - f x <..} \<inter> {..< x + f x}) = S" (is "?S = S")
 | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1337 | by (fastforce simp: dist_real_def) | 
| 71720 | 1338 | moreover have "generate_topology (range lessThan \<union> range greaterThan) ?S" | 
| 1339 | by (force intro: generate_topology.Basis generate_topology_Union generate_topology.Int) | |
| 1340 | ultimately show "generate_topology (range lessThan \<union> range greaterThan) S" | |
| 1341 | by simp | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1342 | next | 
| 63545 | 1343 | fix S :: "real set" | 
| 1344 | assume "generate_topology (range lessThan \<union> range greaterThan) S" | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1345 |     moreover have "\<And>a::real. open {..<a}"
 | 
| 62101 | 1346 | unfolding open_dist dist_real_def | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1347 | proof clarify | 
| 63545 | 1348 | fix x a :: real | 
| 1349 | assume "x < a" | |
| 1350 |       then have "0 < a - x \<and> (\<forall>y. \<bar>y - x\<bar> < a - x \<longrightarrow> y \<in> {..<a})" by auto
 | |
| 1351 |       then show "\<exists>e>0. \<forall>y. \<bar>y - x\<bar> < e \<longrightarrow> y \<in> {..<a}" ..
 | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1352 | qed | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1353 |     moreover have "\<And>a::real. open {a <..}"
 | 
| 62101 | 1354 | unfolding open_dist dist_real_def | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1355 | proof clarify | 
| 63545 | 1356 | fix x a :: real | 
| 1357 | assume "a < x" | |
| 1358 |       then have "0 < x - a \<and> (\<forall>y. \<bar>y - x\<bar> < x - a \<longrightarrow> y \<in> {a<..})" by auto
 | |
| 1359 |       then show "\<exists>e>0. \<forall>y. \<bar>y - x\<bar> < e \<longrightarrow> y \<in> {a<..}" ..
 | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1360 | qed | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1361 | ultimately show "open S" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1362 | by induct auto | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1363 | qed | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1364 | qed | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1365 | |
| 51775 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51774diff
changeset | 1366 | instance real :: linear_continuum_topology .. | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51481diff
changeset | 1367 | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1368 | lemmas open_real_greaterThan = open_greaterThan[where 'a=real] | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1369 | lemmas open_real_lessThan = open_lessThan[where 'a=real] | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1370 | lemmas open_real_greaterThanLessThan = open_greaterThanLessThan[where 'a=real] | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1371 | lemmas closed_real_atMost = closed_atMost[where 'a=real] | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1372 | lemmas closed_real_atLeast = closed_atLeast[where 'a=real] | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1373 | lemmas closed_real_atLeastAtMost = closed_atLeastAtMost[where 'a=real] | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1374 | |
| 70616 
6bc397bc8e8a
explicit instance real::ordered_real_vector before subclass in ordered_euclidean_space
 immler parents: 
70019diff
changeset | 1375 | instance real :: ordered_real_vector | 
| 
6bc397bc8e8a
explicit instance real::ordered_real_vector before subclass in ordered_euclidean_space
 immler parents: 
70019diff
changeset | 1376 | by standard (auto intro: mult_left_mono mult_right_mono) | 
| 
6bc397bc8e8a
explicit instance real::ordered_real_vector before subclass in ordered_euclidean_space
 immler parents: 
70019diff
changeset | 1377 | |
| 63545 | 1378 | |
| 60758 | 1379 | subsection \<open>Extra type constraints\<close> | 
| 31446 | 1380 | |
| 69593 | 1381 | text \<open>Only allow \<^term>\<open>open\<close> in class \<open>topological_space\<close>.\<close> | 
| 60758 | 1382 | setup \<open>Sign.add_const_constraint | 
| 69593 | 1383 | (\<^const_name>\<open>open\<close>, SOME \<^typ>\<open>'a::topological_space set \<Rightarrow> bool\<close>)\<close> | 
| 31492 
5400beeddb55
replace 'topo' with 'open'; add extra type constraint for 'open'
 huffman parents: 
31490diff
changeset | 1384 | |
| 69593 | 1385 | text \<open>Only allow \<^term>\<open>uniformity\<close> in class \<open>uniform_space\<close>.\<close> | 
| 62101 | 1386 | setup \<open>Sign.add_const_constraint | 
| 69593 | 1387 |   (\<^const_name>\<open>uniformity\<close>, SOME \<^typ>\<open>('a::uniformity \<times> 'a) filter\<close>)\<close>
 | 
| 62101 | 1388 | |
| 69593 | 1389 | text \<open>Only allow \<^term>\<open>dist\<close> in class \<open>metric_space\<close>.\<close> | 
| 60758 | 1390 | setup \<open>Sign.add_const_constraint | 
| 69593 | 1391 | (\<^const_name>\<open>dist\<close>, SOME \<^typ>\<open>'a::metric_space \<Rightarrow> 'a \<Rightarrow> real\<close>)\<close> | 
| 31446 | 1392 | |
| 69593 | 1393 | text \<open>Only allow \<^term>\<open>norm\<close> in class \<open>real_normed_vector\<close>.\<close> | 
| 60758 | 1394 | setup \<open>Sign.add_const_constraint | 
| 69593 | 1395 | (\<^const_name>\<open>norm\<close>, SOME \<^typ>\<open>'a::real_normed_vector \<Rightarrow> real\<close>)\<close> | 
| 31446 | 1396 | |
| 63545 | 1397 | |
| 60758 | 1398 | subsection \<open>Sign function\<close> | 
| 22972 
3e96b98d37c6
generalized sgn function to work on any real normed vector space
 huffman parents: 
22942diff
changeset | 1399 | |
| 63545 | 1400 | lemma norm_sgn: "norm (sgn x) = (if x = 0 then 0 else 1)" | 
| 1401 | for x :: "'a::real_normed_vector" | |
| 1402 | by (simp add: sgn_div_norm) | |
| 22972 
3e96b98d37c6
generalized sgn function to work on any real normed vector space
 huffman parents: 
22942diff
changeset | 1403 | |
| 63545 | 1404 | lemma sgn_zero [simp]: "sgn (0::'a::real_normed_vector) = 0" | 
| 1405 | by (simp add: sgn_div_norm) | |
| 22972 
3e96b98d37c6
generalized sgn function to work on any real normed vector space
 huffman parents: 
22942diff
changeset | 1406 | |
| 63545 | 1407 | lemma sgn_zero_iff: "sgn x = 0 \<longleftrightarrow> x = 0" | 
| 1408 | for x :: "'a::real_normed_vector" | |
| 1409 | by (simp add: sgn_div_norm) | |
| 22972 
3e96b98d37c6
generalized sgn function to work on any real normed vector space
 huffman parents: 
22942diff
changeset | 1410 | |
| 63545 | 1411 | lemma sgn_minus: "sgn (- x) = - sgn x" | 
| 1412 | for x :: "'a::real_normed_vector" | |
| 1413 | by (simp add: sgn_div_norm) | |
| 22972 
3e96b98d37c6
generalized sgn function to work on any real normed vector space
 huffman parents: 
22942diff
changeset | 1414 | |
| 63545 | 1415 | lemma sgn_scaleR: "sgn (scaleR r x) = scaleR (sgn r) (sgn x)" | 
| 1416 | for x :: "'a::real_normed_vector" | |
| 1417 | by (simp add: sgn_div_norm ac_simps) | |
| 22973 
64d300e16370
add lemma sgn_mult; declare real_scaleR_def and scaleR_eq_0_iff as simp rules
 huffman parents: 
22972diff
changeset | 1418 | |
| 22972 
3e96b98d37c6
generalized sgn function to work on any real normed vector space
 huffman parents: 
22942diff
changeset | 1419 | lemma sgn_one [simp]: "sgn (1::'a::real_normed_algebra_1) = 1" | 
| 63545 | 1420 | by (simp add: sgn_div_norm) | 
| 22972 
3e96b98d37c6
generalized sgn function to work on any real normed vector space
 huffman parents: 
22942diff
changeset | 1421 | |
| 63545 | 1422 | lemma sgn_of_real: "sgn (of_real r :: 'a::real_normed_algebra_1) = of_real (sgn r)" | 
| 1423 | unfolding of_real_def by (simp only: sgn_scaleR sgn_one) | |
| 22972 
3e96b98d37c6
generalized sgn function to work on any real normed vector space
 huffman parents: 
22942diff
changeset | 1424 | |
| 63545 | 1425 | lemma sgn_mult: "sgn (x * y) = sgn x * sgn y" | 
| 1426 | for x y :: "'a::real_normed_div_algebra" | |
| 71544 | 1427 | by (simp add: sgn_div_norm norm_mult) | 
| 22973 
64d300e16370
add lemma sgn_mult; declare real_scaleR_def and scaleR_eq_0_iff as simp rules
 huffman parents: 
22972diff
changeset | 1428 | |
| 64240 | 1429 | hide_fact (open) sgn_mult | 
| 1430 | ||
| 63545 | 1431 | lemma real_sgn_eq: "sgn x = x / \<bar>x\<bar>" | 
| 1432 | for x :: real | |
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 1433 | by (simp add: sgn_div_norm divide_inverse) | 
| 22972 
3e96b98d37c6
generalized sgn function to work on any real normed vector space
 huffman parents: 
22942diff
changeset | 1434 | |
| 63545 | 1435 | lemma zero_le_sgn_iff [simp]: "0 \<le> sgn x \<longleftrightarrow> 0 \<le> x" | 
| 1436 | for x :: real | |
| 56889 
48a745e1bde7
avoid the Complex constructor, use the more natural Re/Im view; moved csqrt to Complex.
 hoelzl parents: 
56479diff
changeset | 1437 | by (cases "0::real" x rule: linorder_cases) simp_all | 
| 60026 
41d81b4a0a21
Restored LIMSEQ_def as legacy binding. [The other changes are whitespace only.]
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 1438 | |
| 63545 | 1439 | lemma sgn_le_0_iff [simp]: "sgn x \<le> 0 \<longleftrightarrow> x \<le> 0" | 
| 1440 | for x :: real | |
| 56889 
48a745e1bde7
avoid the Complex constructor, use the more natural Re/Im view; moved csqrt to Complex.
 hoelzl parents: 
56479diff
changeset | 1441 | by (cases "0::real" x rule: linorder_cases) simp_all | 
| 60026 
41d81b4a0a21
Restored LIMSEQ_def as legacy binding. [The other changes are whitespace only.]
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 1442 | |
| 51474 
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
 hoelzl parents: 
51472diff
changeset | 1443 | lemma norm_conv_dist: "norm x = dist x 0" | 
| 
1e9e68247ad1
generalize Bfun and Bseq to metric spaces; Bseq is an abbreviation for Bfun
 hoelzl parents: 
51472diff
changeset | 1444 | unfolding dist_norm by simp | 
| 22972 
3e96b98d37c6
generalized sgn function to work on any real normed vector space
 huffman parents: 
22942diff
changeset | 1445 | |
| 62379 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62368diff
changeset | 1446 | declare norm_conv_dist [symmetric, simp] | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62368diff
changeset | 1447 | |
| 63545 | 1448 | lemma dist_0_norm [simp]: "dist 0 x = norm x" | 
| 1449 | for x :: "'a::real_normed_vector" | |
| 1450 | by (simp add: dist_norm) | |
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62379diff
changeset | 1451 | |
| 60307 
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
 paulson <lp15@cam.ac.uk> parents: 
60303diff
changeset | 1452 | lemma dist_diff [simp]: "dist a (a - b) = norm b" "dist (a - b) a = norm b" | 
| 
75e1aa7a450e
Convex hulls: theorems about interior, etc. And a few simple lemmas.
 paulson <lp15@cam.ac.uk> parents: 
60303diff
changeset | 1453 | by (simp_all add: dist_norm) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 1454 | |
| 61524 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61169diff
changeset | 1455 | lemma dist_of_int: "dist (of_int m) (of_int n :: 'a :: real_normed_algebra_1) = of_int \<bar>m - n\<bar>" | 
| 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61169diff
changeset | 1456 | proof - | 
| 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61169diff
changeset | 1457 | have "dist (of_int m) (of_int n :: 'a) = dist (of_int m :: 'a) (of_int m - (of_int (m - n)))" | 
| 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61169diff
changeset | 1458 | by simp | 
| 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61169diff
changeset | 1459 | also have "\<dots> = of_int \<bar>m - n\<bar>" by (subst dist_diff, subst norm_of_int) simp | 
| 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61169diff
changeset | 1460 | finally show ?thesis . | 
| 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61169diff
changeset | 1461 | qed | 
| 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61169diff
changeset | 1462 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 1463 | lemma dist_of_nat: | 
| 61524 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61169diff
changeset | 1464 | "dist (of_nat m) (of_nat n :: 'a :: real_normed_algebra_1) = of_int \<bar>int m - int n\<bar>" | 
| 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61169diff
changeset | 1465 | by (subst (1 2) of_int_of_nat_eq [symmetric]) (rule dist_of_int) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 1466 | |
| 63545 | 1467 | |
| 60758 | 1468 | subsection \<open>Bounded Linear and Bilinear Operators\<close> | 
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1469 | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 1470 | lemma linearI: "linear f" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 1471 | if "\<And>b1 b2. f (b1 + b2) = f b1 + f b2" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 1472 | "\<And>r b. f (r *\<^sub>R b) = r *\<^sub>R f b" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 1473 | using that | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 1474 | by unfold_locales (auto simp: algebra_simps) | 
| 53600 
8fda7ad57466
make 'linear' into a sublocale of 'bounded_linear';
 huffman parents: 
53381diff
changeset | 1475 | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 1476 | lemma linear_iff: | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 1477 | "linear f \<longleftrightarrow> (\<forall>x y. f (x + y) = f x + f y) \<and> (\<forall>c x. f (c *\<^sub>R x) = c *\<^sub>R f x)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 1478 | (is "linear f \<longleftrightarrow> ?rhs") | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 1479 | proof | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 1480 | assume "linear f" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 1481 | then interpret f: linear f . | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 1482 | show "?rhs" by (simp add: f.add f.scale) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 1483 | next | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 1484 | assume "?rhs" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 1485 | then show "linear f" by (intro linearI) auto | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 1486 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 1487 | |
| 79945 
ca004ccf2352
New material from a variety of sources (including AFP)
 paulson <lp15@cam.ac.uk> parents: 
78656diff
changeset | 1488 | lemma linear_of_real [simp]: "linear of_real" | 
| 
ca004ccf2352
New material from a variety of sources (including AFP)
 paulson <lp15@cam.ac.uk> parents: 
78656diff
changeset | 1489 | by (simp add: linear_iff scaleR_conv_of_real) | 
| 
ca004ccf2352
New material from a variety of sources (including AFP)
 paulson <lp15@cam.ac.uk> parents: 
78656diff
changeset | 1490 | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 1491 | lemmas linear_scaleR_left = linear_scale_left | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 1492 | lemmas linear_imp_scaleR = linear_imp_scale | 
| 60800 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 1493 | |
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 1494 | corollary real_linearD: | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 1495 | fixes f :: "real \<Rightarrow> real" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68721diff
changeset | 1496 | assumes "linear f" obtains c where "f = (*) c" | 
| 63545 | 1497 | by (rule linear_imp_scaleR [OF assms]) (force simp: scaleR_conv_of_real) | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 1498 | |
| 65583 
8d53b3bebab4
Further new material. The simprule status of some exp and ln identities was reverted.
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 1499 | lemma linear_times_of_real: "linear (\<lambda>x. a * of_real x)" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 1500 | by (auto intro!: linearI simp: distrib_left) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67727diff
changeset | 1501 | (metis mult_scaleR_right scaleR_conv_of_real) | 
| 53600 
8fda7ad57466
make 'linear' into a sublocale of 'bounded_linear';
 huffman parents: 
53381diff
changeset | 1502 | |
| 
8fda7ad57466
make 'linear' into a sublocale of 'bounded_linear';
 huffman parents: 
53381diff
changeset | 1503 | locale bounded_linear = linear f for f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" + | 
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1504 | assumes bounded: "\<exists>K. \<forall>x. norm (f x) \<le> norm x * K" | 
| 27443 | 1505 | begin | 
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1506 | |
| 63545 | 1507 | lemma pos_bounded: "\<exists>K>0. \<forall>x. norm (f x) \<le> norm x * K" | 
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1508 | proof - | 
| 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1509 | obtain K where K: "\<And>x. norm (f x) \<le> norm x * K" | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 1510 | using bounded by blast | 
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1511 | show ?thesis | 
| 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1512 | proof (intro exI impI conjI allI) | 
| 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1513 | show "0 < max 1 K" | 
| 54863 
82acc20ded73
prefer more canonical names for lemmas on min/max
 haftmann parents: 
54785diff
changeset | 1514 | by (rule order_less_le_trans [OF zero_less_one max.cobounded1]) | 
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1515 | next | 
| 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1516 | fix x | 
| 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1517 | have "norm (f x) \<le> norm x * K" using K . | 
| 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1518 | also have "\<dots> \<le> norm x * max 1 K" | 
| 54863 
82acc20ded73
prefer more canonical names for lemmas on min/max
 haftmann parents: 
54785diff
changeset | 1519 | by (rule mult_left_mono [OF max.cobounded2 norm_ge_zero]) | 
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1520 | finally show "norm (f x) \<le> norm x * max 1 K" . | 
| 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1521 | qed | 
| 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1522 | qed | 
| 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1523 | |
| 63545 | 1524 | lemma nonneg_bounded: "\<exists>K\<ge>0. \<forall>x. norm (f x) \<le> norm x * K" | 
| 1525 | using pos_bounded by (auto intro: order_less_imp_le) | |
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1526 | |
| 63545 | 1527 | lemma linear: "linear f" | 
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63128diff
changeset | 1528 | by (fact local.linear_axioms) | 
| 56369 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56194diff
changeset | 1529 | |
| 27443 | 1530 | end | 
| 1531 | ||
| 44127 | 1532 | lemma bounded_linear_intro: | 
| 1533 | assumes "\<And>x y. f (x + y) = f x + f y" | |
| 63545 | 1534 | and "\<And>r x. f (scaleR r x) = scaleR r (f x)" | 
| 1535 | and "\<And>x. norm (f x) \<le> norm x * K" | |
| 44127 | 1536 | shows "bounded_linear f" | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 1537 | by standard (blast intro: assms)+ | 
| 44127 | 1538 | |
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1539 | locale bounded_bilinear = | 
| 63545 | 1540 | fixes prod :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector \<Rightarrow> 'c::real_normed_vector" | 
| 80932 
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
 wenzelm parents: 
79945diff
changeset | 1541 | (infixl \<open>**\<close> 70) | 
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1542 | assumes add_left: "prod (a + a') b = prod a b + prod a' b" | 
| 63545 | 1543 | and add_right: "prod a (b + b') = prod a b + prod a b'" | 
| 1544 | and scaleR_left: "prod (scaleR r a) b = scaleR r (prod a b)" | |
| 1545 | and scaleR_right: "prod a (scaleR r b) = scaleR r (prod a b)" | |
| 1546 | and bounded: "\<exists>K. \<forall>a b. norm (prod a b) \<le> norm a * norm b * K" | |
| 27443 | 1547 | begin | 
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1548 | |
| 63545 | 1549 | lemma pos_bounded: "\<exists>K>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K" | 
| 66793 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 1550 | proof - | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 1551 | obtain K where "\<And>a b. norm (a ** b) \<le> norm a * norm b * K" | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 1552 | using bounded by blast | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 1553 | then have "norm (a ** b) \<le> norm a * norm b * (max 1 K)" for a b | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 1554 | by (rule order.trans) (simp add: mult_left_mono) | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 1555 | then show ?thesis | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 1556 | by force | 
| 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 1557 | qed | 
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1558 | |
| 63545 | 1559 | lemma nonneg_bounded: "\<exists>K\<ge>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K" | 
| 1560 | using pos_bounded by (auto intro: order_less_imp_le) | |
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1561 | |
| 27443 | 1562 | lemma additive_right: "additive (\<lambda>b. prod a b)" | 
| 63545 | 1563 | by (rule additive.intro, rule add_right) | 
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1564 | |
| 27443 | 1565 | lemma additive_left: "additive (\<lambda>a. prod a b)" | 
| 63545 | 1566 | by (rule additive.intro, rule add_left) | 
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1567 | |
| 27443 | 1568 | lemma zero_left: "prod 0 b = 0" | 
| 63545 | 1569 | by (rule additive.zero [OF additive_left]) | 
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1570 | |
| 27443 | 1571 | lemma zero_right: "prod a 0 = 0" | 
| 63545 | 1572 | by (rule additive.zero [OF additive_right]) | 
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1573 | |
| 27443 | 1574 | lemma minus_left: "prod (- a) b = - prod a b" | 
| 63545 | 1575 | by (rule additive.minus [OF additive_left]) | 
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1576 | |
| 27443 | 1577 | lemma minus_right: "prod a (- b) = - prod a b" | 
| 63545 | 1578 | by (rule additive.minus [OF additive_right]) | 
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1579 | |
| 63545 | 1580 | lemma diff_left: "prod (a - a') b = prod a b - prod a' b" | 
| 1581 | by (rule additive.diff [OF additive_left]) | |
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1582 | |
| 63545 | 1583 | lemma diff_right: "prod a (b - b') = prod a b - prod a b'" | 
| 1584 | by (rule additive.diff [OF additive_right]) | |
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1585 | |
| 64267 | 1586 | lemma sum_left: "prod (sum g S) x = sum ((\<lambda>i. prod (g i) x)) S" | 
| 1587 | by (rule additive.sum [OF additive_left]) | |
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61799diff
changeset | 1588 | |
| 64267 | 1589 | lemma sum_right: "prod x (sum g S) = sum ((\<lambda>i. (prod x (g i)))) S" | 
| 1590 | by (rule additive.sum [OF additive_right]) | |
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61799diff
changeset | 1591 | |
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61799diff
changeset | 1592 | |
| 63545 | 1593 | lemma bounded_linear_left: "bounded_linear (\<lambda>a. a ** b)" | 
| 68594 | 1594 | proof - | 
| 1595 | obtain K where "\<And>a b. norm (a ** b) \<le> norm a * norm b * K" | |
| 1596 | using pos_bounded by blast | |
| 1597 | then show ?thesis | |
| 1598 | by (rule_tac K="norm b * K" in bounded_linear_intro) (auto simp: algebra_simps scaleR_left add_left) | |
| 1599 | qed | |
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1600 | |
| 63545 | 1601 | lemma bounded_linear_right: "bounded_linear (\<lambda>b. a ** b)" | 
| 68594 | 1602 | proof - | 
| 1603 | obtain K where "\<And>a b. norm (a ** b) \<le> norm a * norm b * K" | |
| 1604 | using pos_bounded by blast | |
| 1605 | then show ?thesis | |
| 1606 | by (rule_tac K="norm a * K" in bounded_linear_intro) (auto simp: algebra_simps scaleR_right add_right) | |
| 1607 | qed | |
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1608 | |
| 63545 | 1609 | lemma prod_diff_prod: "(x ** y - a ** b) = (x - a) ** (y - b) + (x - a) ** b + a ** (y - b)" | 
| 1610 | by (simp add: diff_left diff_right) | |
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1611 | |
| 61916 | 1612 | lemma flip: "bounded_bilinear (\<lambda>x y. y ** x)" | 
| 71720 | 1613 | proof | 
| 1614 | show "\<exists>K. \<forall>a b. norm (b ** a) \<le> norm a * norm b * K" | |
| 1615 | by (metis bounded mult.commute) | |
| 1616 | qed (simp_all add: add_right add_left scaleR_right scaleR_left) | |
| 61916 | 1617 | |
| 1618 | lemma comp1: | |
| 1619 | assumes "bounded_linear g" | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68721diff
changeset | 1620 | shows "bounded_bilinear (\<lambda>x. (**) (g x))" | 
| 61916 | 1621 | proof unfold_locales | 
| 1622 | interpret g: bounded_linear g by fact | |
| 1623 | show "\<And>a a' b. g (a + a') ** b = g a ** b + g a' ** b" | |
| 1624 | "\<And>a b b'. g a ** (b + b') = g a ** b + g a ** b'" | |
| 1625 | "\<And>r a b. g (r *\<^sub>R a) ** b = r *\<^sub>R (g a ** b)" | |
| 1626 | "\<And>a r b. g a ** (r *\<^sub>R b) = r *\<^sub>R (g a ** b)" | |
| 1627 | by (auto simp: g.add add_left add_right g.scaleR scaleR_left scaleR_right) | |
| 63545 | 1628 | from g.nonneg_bounded nonneg_bounded obtain K L | 
| 1629 | where nn: "0 \<le> K" "0 \<le> L" | |
| 1630 | and K: "\<And>x. norm (g x) \<le> norm x * K" | |
| 1631 | and L: "\<And>a b. norm (a ** b) \<le> norm a * norm b * L" | |
| 61916 | 1632 | by auto | 
| 1633 | have "norm (g a ** b) \<le> norm a * K * norm b * L" for a b | |
| 1634 | by (auto intro!: order_trans[OF K] order_trans[OF L] mult_mono simp: nn) | |
| 1635 | then show "\<exists>K. \<forall>a b. norm (g a ** b) \<le> norm a * norm b * K" | |
| 1636 | by (auto intro!: exI[where x="K * L"] simp: ac_simps) | |
| 1637 | qed | |
| 1638 | ||
| 63545 | 1639 | lemma comp: "bounded_linear f \<Longrightarrow> bounded_linear g \<Longrightarrow> bounded_bilinear (\<lambda>x y. f x ** g y)" | 
| 61916 | 1640 | by (rule bounded_bilinear.flip[OF bounded_bilinear.comp1[OF bounded_bilinear.flip[OF comp1]]]) | 
| 1641 | ||
| 27443 | 1642 | end | 
| 1643 | ||
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1644 | lemma bounded_linear_ident[simp]: "bounded_linear (\<lambda>x. x)" | 
| 61169 | 1645 | by standard (auto intro!: exI[of _ 1]) | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1646 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1647 | lemma bounded_linear_zero[simp]: "bounded_linear (\<lambda>x. 0)" | 
| 61169 | 1648 | by standard (auto intro!: exI[of _ 1]) | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1649 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1650 | lemma bounded_linear_add: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1651 | assumes "bounded_linear f" | 
| 63545 | 1652 | and "bounded_linear g" | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1653 | shows "bounded_linear (\<lambda>x. f x + g x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1654 | proof - | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1655 | interpret f: bounded_linear f by fact | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1656 | interpret g: bounded_linear g by fact | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1657 | show ?thesis | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1658 | proof | 
| 63545 | 1659 | from f.bounded obtain Kf where Kf: "norm (f x) \<le> norm x * Kf" for x | 
| 1660 | by blast | |
| 1661 | from g.bounded obtain Kg where Kg: "norm (g x) \<le> norm x * Kg" for x | |
| 1662 | by blast | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1663 | show "\<exists>K. \<forall>x. norm (f x + g x) \<le> norm x * K" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1664 | using add_mono[OF Kf Kg] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1665 | by (intro exI[of _ "Kf + Kg"]) (auto simp: field_simps intro: norm_triangle_ineq order_trans) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1666 | qed (simp_all add: f.add g.add f.scaleR g.scaleR scaleR_right_distrib) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1667 | qed | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1668 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1669 | lemma bounded_linear_minus: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1670 | assumes "bounded_linear f" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1671 | shows "bounded_linear (\<lambda>x. - f x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1672 | proof - | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1673 | interpret f: bounded_linear f by fact | 
| 63545 | 1674 | show ?thesis | 
| 68669 | 1675 | by unfold_locales (simp_all add: f.add f.scaleR f.bounded) | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1676 | qed | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1677 | |
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61799diff
changeset | 1678 | lemma bounded_linear_sub: "bounded_linear f \<Longrightarrow> bounded_linear g \<Longrightarrow> bounded_linear (\<lambda>x. f x - g x)" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61799diff
changeset | 1679 | using bounded_linear_add[of f "\<lambda>x. - g x"] bounded_linear_minus[of g] | 
| 68594 | 1680 | by (auto simp: algebra_simps) | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61799diff
changeset | 1681 | |
| 64267 | 1682 | lemma bounded_linear_sum: | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61799diff
changeset | 1683 | fixes f :: "'i \<Rightarrow> 'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 63915 | 1684 | shows "(\<And>i. i \<in> I \<Longrightarrow> bounded_linear (f i)) \<Longrightarrow> bounded_linear (\<lambda>x. \<Sum>i\<in>I. f i x)" | 
| 1685 | by (induct I rule: infinite_finite_induct) (auto intro!: bounded_linear_add) | |
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61799diff
changeset | 1686 | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1687 | lemma bounded_linear_compose: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1688 | assumes "bounded_linear f" | 
| 63545 | 1689 | and "bounded_linear g" | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1690 | shows "bounded_linear (\<lambda>x. f (g x))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1691 | proof - | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1692 | interpret f: bounded_linear f by fact | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1693 | interpret g: bounded_linear g by fact | 
| 63545 | 1694 | show ?thesis | 
| 1695 | proof unfold_locales | |
| 1696 | show "f (g (x + y)) = f (g x) + f (g y)" for x y | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1697 | by (simp only: f.add g.add) | 
| 63545 | 1698 | show "f (g (scaleR r x)) = scaleR r (f (g x))" for r x | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1699 | by (simp only: f.scaleR g.scaleR) | 
| 63545 | 1700 | from f.pos_bounded obtain Kf where f: "\<And>x. norm (f x) \<le> norm x * Kf" and Kf: "0 < Kf" | 
| 1701 | by blast | |
| 1702 | from g.pos_bounded obtain Kg where g: "\<And>x. norm (g x) \<le> norm x * Kg" | |
| 1703 | by blast | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1704 | show "\<exists>K. \<forall>x. norm (f (g x)) \<le> norm x * K" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1705 | proof (intro exI allI) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1706 | fix x | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1707 | have "norm (f (g x)) \<le> norm (g x) * Kf" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1708 | using f . | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1709 | also have "\<dots> \<le> (norm x * Kg) * Kf" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1710 | using g Kf [THEN order_less_imp_le] by (rule mult_right_mono) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1711 | also have "(norm x * Kg) * Kf = norm x * (Kg * Kf)" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57448diff
changeset | 1712 | by (rule mult.assoc) | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1713 | finally show "norm (f (g x)) \<le> norm x * (Kg * Kf)" . | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1714 | qed | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1715 | qed | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1716 | qed | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1717 | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68721diff
changeset | 1718 | lemma bounded_bilinear_mult: "bounded_bilinear ((*) :: 'a \<Rightarrow> 'a \<Rightarrow> 'a::real_normed_algebra)" | 
| 71720 | 1719 | proof (rule bounded_bilinear.intro) | 
| 1720 | show "\<exists>K. \<forall>a b::'a. norm (a * b) \<le> norm a * norm b * K" | |
| 1721 | by (rule_tac x=1 in exI) (simp add: norm_mult_ineq) | |
| 1722 | qed (auto simp: algebra_simps) | |
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1723 | |
| 63545 | 1724 | lemma bounded_linear_mult_left: "bounded_linear (\<lambda>x::'a::real_normed_algebra. x * y)" | 
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44127diff
changeset | 1725 | using bounded_bilinear_mult | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44127diff
changeset | 1726 | by (rule bounded_bilinear.bounded_linear_left) | 
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1727 | |
| 63545 | 1728 | lemma bounded_linear_mult_right: "bounded_linear (\<lambda>y::'a::real_normed_algebra. x * y)" | 
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44127diff
changeset | 1729 | using bounded_bilinear_mult | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44127diff
changeset | 1730 | by (rule bounded_bilinear.bounded_linear_right) | 
| 23127 | 1731 | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1732 | lemmas bounded_linear_mult_const = | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1733 | bounded_linear_mult_left [THEN bounded_linear_compose] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1734 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1735 | lemmas bounded_linear_const_mult = | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1736 | bounded_linear_mult_right [THEN bounded_linear_compose] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1737 | |
| 63545 | 1738 | lemma bounded_linear_divide: "bounded_linear (\<lambda>x. x / y)" | 
| 1739 | for y :: "'a::real_normed_field" | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44127diff
changeset | 1740 | unfolding divide_inverse by (rule bounded_linear_mult_left) | 
| 23120 | 1741 | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44127diff
changeset | 1742 | lemma bounded_bilinear_scaleR: "bounded_bilinear scaleR" | 
| 71720 | 1743 | proof (rule bounded_bilinear.intro) | 
| 1744 | show "\<exists>K. \<forall>a b. norm (a *\<^sub>R b) \<le> norm a * norm b * K" | |
| 1745 | using less_eq_real_def by auto | |
| 1746 | qed (auto simp: algebra_simps) | |
| 22442 
15d9ed9b5051
move bounded (bi)linear operator locales from Lim.thy to RealVector.thy
 huffman parents: 
21809diff
changeset | 1747 | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44127diff
changeset | 1748 | lemma bounded_linear_scaleR_left: "bounded_linear (\<lambda>r. scaleR r x)" | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44127diff
changeset | 1749 | using bounded_bilinear_scaleR | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44127diff
changeset | 1750 | by (rule bounded_bilinear.bounded_linear_left) | 
| 23127 | 1751 | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44127diff
changeset | 1752 | lemma bounded_linear_scaleR_right: "bounded_linear (\<lambda>x. scaleR r x)" | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44127diff
changeset | 1753 | using bounded_bilinear_scaleR | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44127diff
changeset | 1754 | by (rule bounded_bilinear.bounded_linear_right) | 
| 23127 | 1755 | |
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61799diff
changeset | 1756 | lemmas bounded_linear_scaleR_const = | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61799diff
changeset | 1757 | bounded_linear_scaleR_left[THEN bounded_linear_compose] | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61799diff
changeset | 1758 | |
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61799diff
changeset | 1759 | lemmas bounded_linear_const_scaleR = | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61799diff
changeset | 1760 | bounded_linear_scaleR_right[THEN bounded_linear_compose] | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61799diff
changeset | 1761 | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44127diff
changeset | 1762 | lemma bounded_linear_of_real: "bounded_linear (\<lambda>r. of_real r)" | 
| 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44127diff
changeset | 1763 | unfolding of_real_def by (rule bounded_linear_scaleR_left) | 
| 22625 | 1764 | |
| 63545 | 1765 | lemma real_bounded_linear: "bounded_linear f \<longleftrightarrow> (\<exists>c::real. f = (\<lambda>x. x * c))" | 
| 1766 | for f :: "real \<Rightarrow> real" | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1767 | proof - | 
| 63545 | 1768 |   {
 | 
| 1769 | fix x | |
| 1770 | assume "bounded_linear f" | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1771 | then interpret bounded_linear f . | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1772 | from scaleR[of x 1] have "f x = x * f 1" | 
| 63545 | 1773 | by simp | 
| 1774 | } | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1775 | then show ?thesis | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1776 | by (auto intro: exI[of _ "f 1"] bounded_linear_mult_left) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1777 | qed | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1778 | |
| 44571 | 1779 | instance real_normed_algebra_1 \<subseteq> perfect_space | 
| 1780 | proof | |
| 71720 | 1781 | fix x::'a | 
| 1782 | have "\<And>e. 0 < e \<Longrightarrow> \<exists>y. norm (y - x) < e \<and> y \<noteq> x" | |
| 1783 | by (rule_tac x = "x + of_real (e/2)" in exI) auto | |
| 1784 |   then show "\<not> open {x}" 
 | |
| 1785 | by (clarsimp simp: open_dist dist_norm) | |
| 44571 | 1786 | qed | 
| 1787 | ||
| 63545 | 1788 | |
| 60758 | 1789 | subsection \<open>Filters and Limits on Metric Space\<close> | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1790 | |
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69064diff
changeset | 1791 | lemma (in metric_space) nhds_metric: "nhds x = (INF e\<in>{0 <..}. principal {y. dist y x < e})"
 | 
| 57448 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 1792 | unfolding nhds_def | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 1793 | proof (safe intro!: INF_eq) | 
| 63545 | 1794 | fix S | 
| 1795 | assume "open S" "x \<in> S" | |
| 57448 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 1796 |   then obtain e where "{y. dist y x < e} \<subseteq> S" "0 < e"
 | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 1797 | by (auto simp: open_dist subset_eq) | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 1798 |   then show "\<exists>e\<in>{0<..}. principal {y. dist y x < e} \<le> principal S"
 | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 1799 | by auto | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 1800 | qed (auto intro!: exI[of _ "{y. dist x y < e}" for e] open_ball simp: dist_commute)
 | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 1801 | |
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1802 | (* Contributed by Dominique Unruh *) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1803 | lemma tendsto_iff_uniformity: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1804 | \<comment> \<open>More general analogus of \<open>tendsto_iff\<close> below. Applies to all uniform spaces, not just metric ones.\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1805 | fixes l :: \<open>'b :: uniform_space\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1806 | shows \<open>(f \<longlongrightarrow> l) F \<longleftrightarrow> (\<forall>E. eventually E uniformity \<longrightarrow> (\<forall>\<^sub>F x in F. E (f x, l)))\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1807 | proof (intro iffI allI impI) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1808 |   fix E :: \<open>('b \<times> 'b) \<Rightarrow> bool\<close>
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1809 | assume \<open>(f \<longlongrightarrow> l) F\<close> and \<open>eventually E uniformity\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1810 | from \<open>eventually E uniformity\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1811 | have \<open>eventually (\<lambda>(x, y). E (y, x)) uniformity\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1812 | by (simp add: uniformity_sym) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1813 | then have \<open>\<forall>\<^sub>F (y, x) in uniformity. y = l \<longrightarrow> E (x, y)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1814 | using eventually_mono by fastforce | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1815 | with \<open>(f \<longlongrightarrow> l) F\<close> have \<open>eventually (\<lambda>x. E (x ,l)) (filtermap f F)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1816 | by (simp add: filterlim_def le_filter_def eventually_nhds_uniformity) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1817 | then show \<open>\<forall>\<^sub>F x in F. E (f x, l)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1818 | by (simp add: eventually_filtermap) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1819 | next | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1820 | assume assm: \<open>\<forall>E. eventually E uniformity \<longrightarrow> (\<forall>\<^sub>F x in F. E (f x, l))\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1821 | have \<open>eventually P (filtermap f F)\<close> if \<open>\<forall>\<^sub>F (x, y) in uniformity. x = l \<longrightarrow> P y\<close> for P | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1822 | proof - | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1823 | from that have \<open>\<forall>\<^sub>F (y, x) in uniformity. x = l \<longrightarrow> P y\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1824 | using uniformity_sym[where E=\<open>\<lambda>(x,y). x=l \<longrightarrow> P y\<close>] by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1825 | with assm have \<open>\<forall>\<^sub>F x in F. P (f x)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1826 | by auto | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1827 | then show ?thesis | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1828 | by (auto simp: eventually_filtermap) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1829 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1830 | then show \<open>(f \<longlongrightarrow> l) F\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1831 | by (simp add: filterlim_def le_filter_def eventually_nhds_uniformity) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1832 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 1833 | |
| 63545 | 1834 | lemma (in metric_space) tendsto_iff: "(f \<longlongrightarrow> l) F \<longleftrightarrow> (\<forall>e>0. eventually (\<lambda>x. dist (f x) l < e) F)" | 
| 57448 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 1835 | unfolding nhds_metric filterlim_INF filterlim_principal by auto | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 1836 | |
| 67727 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 1837 | lemma tendsto_dist_iff: | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 1838 | "((f \<longlongrightarrow> l) F) \<longleftrightarrow> (((\<lambda>x. dist (f x) l) \<longlongrightarrow> 0) F)" | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 1839 | unfolding tendsto_iff by simp | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 1840 | |
| 63545 | 1841 | lemma (in metric_space) tendstoI [intro?]: | 
| 1842 | "(\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>x. dist (f x) l < e) F) \<Longrightarrow> (f \<longlongrightarrow> l) F" | |
| 57448 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 1843 | by (auto simp: tendsto_iff) | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 1844 | |
| 61973 | 1845 | lemma (in metric_space) tendstoD: "(f \<longlongrightarrow> l) F \<Longrightarrow> 0 < e \<Longrightarrow> eventually (\<lambda>x. dist (f x) l < e) F" | 
| 57448 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 1846 | by (auto simp: tendsto_iff) | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 1847 | |
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 1848 | lemma (in metric_space) eventually_nhds_metric: | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 1849 | "eventually P (nhds a) \<longleftrightarrow> (\<exists>d>0. \<forall>x. dist x a < d \<longrightarrow> P x)" | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 1850 | unfolding nhds_metric | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 1851 | by (subst eventually_INF_base) | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 1852 | (auto simp: eventually_principal Bex_def subset_eq intro: exI[of _ "min a b" for a b]) | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1853 | |
| 63545 | 1854 | lemma eventually_at: "eventually P (at a within S) \<longleftrightarrow> (\<exists>d>0. \<forall>x\<in>S. x \<noteq> a \<and> dist x a < d \<longrightarrow> P x)" | 
| 1855 | for a :: "'a :: metric_space" | |
| 1856 | by (auto simp: eventually_at_filter eventually_nhds_metric) | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1857 | |
| 67706 
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
 Wenda Li <wl302@cam.ac.uk> parents: 
67673diff
changeset | 1858 | lemma frequently_at: "frequently P (at a within S) \<longleftrightarrow> (\<forall>d>0. \<exists>x\<in>S. x \<noteq> a \<and> dist x a < d \<and> P x)" | 
| 
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
 Wenda Li <wl302@cam.ac.uk> parents: 
67673diff
changeset | 1859 | for a :: "'a :: metric_space" | 
| 
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
 Wenda Li <wl302@cam.ac.uk> parents: 
67673diff
changeset | 1860 | unfolding frequently_def eventually_at by auto | 
| 
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
 Wenda Li <wl302@cam.ac.uk> parents: 
67673diff
changeset | 1861 | |
| 63545 | 1862 | lemma eventually_at_le: "eventually P (at a within S) \<longleftrightarrow> (\<exists>d>0. \<forall>x\<in>S. x \<noteq> a \<and> dist x a \<le> d \<longrightarrow> P x)" | 
| 1863 | for a :: "'a::metric_space" | |
| 68594 | 1864 | unfolding eventually_at_filter eventually_nhds_metric | 
| 1865 | apply safe | |
| 1866 | apply (rule_tac x="d / 2" in exI, auto) | |
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51531diff
changeset | 1867 | done | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1868 | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1869 | lemma eventually_at_left_real: "a > (b :: real) \<Longrightarrow> eventually (\<lambda>x. x \<in> {b<..<a}) (at_left a)"
 | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1870 | by (subst eventually_at, rule exI[of _ "a - b"]) (force simp: dist_real_def) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1871 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1872 | lemma eventually_at_right_real: "a < (b :: real) \<Longrightarrow> eventually (\<lambda>x. x \<in> {a<..<b}) (at_right a)"
 | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1873 | by (subst eventually_at, rule exI[of _ "b - a"]) (force simp: dist_real_def) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 1874 | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1875 | lemma metric_tendsto_imp_tendsto: | 
| 63545 | 1876 | fixes a :: "'a :: metric_space" | 
| 1877 | and b :: "'b :: metric_space" | |
| 61973 | 1878 | assumes f: "(f \<longlongrightarrow> a) F" | 
| 63545 | 1879 | and le: "eventually (\<lambda>x. dist (g x) b \<le> dist (f x) a) F" | 
| 61973 | 1880 | shows "(g \<longlongrightarrow> b) F" | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1881 | proof (rule tendstoI) | 
| 63545 | 1882 | fix e :: real | 
| 1883 | assume "0 < e" | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1884 | with f have "eventually (\<lambda>x. dist (f x) a < e) F" by (rule tendstoD) | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1885 | with le show "eventually (\<lambda>x. dist (g x) b < e) F" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1886 | using le_less_trans by (rule eventually_elim2) | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1887 | qed | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1888 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1889 | lemma filterlim_real_sequentially: "LIM x sequentially. real x :> at_top" | 
| 71720 | 1890 | proof (clarsimp simp: filterlim_at_top) | 
| 1891 | fix Z | |
| 1892 | show "\<forall>\<^sub>F x in sequentially. Z \<le> real x" | |
| 1893 | by (meson eventually_sequentiallyI nat_ceiling_le_eq) | |
| 1894 | qed | |
| 61942 | 1895 | |
| 63556 | 1896 | lemma filterlim_nat_sequentially: "filterlim nat sequentially at_top" | 
| 68594 | 1897 | proof - | 
| 1898 | have "\<forall>\<^sub>F x in at_top. Z \<le> nat x" for Z | |
| 1899 | by (auto intro!: eventually_at_top_linorderI[where c="int Z"]) | |
| 1900 | then show ?thesis | |
| 1901 | unfolding filterlim_at_top .. | |
| 1902 | qed | |
| 63556 | 1903 | |
| 1904 | lemma filterlim_floor_sequentially: "filterlim floor at_top at_top" | |
| 68594 | 1905 | proof - | 
| 1906 | have "\<forall>\<^sub>F x in at_top. Z \<le> \<lfloor>x\<rfloor>" for Z | |
| 1907 | by (auto simp: le_floor_iff intro!: eventually_at_top_linorderI[where c="of_int Z"]) | |
| 1908 | then show ?thesis | |
| 1909 | unfolding filterlim_at_top .. | |
| 1910 | qed | |
| 63556 | 1911 | |
| 1912 | lemma filterlim_sequentially_iff_filterlim_real: | |
| 71720 | 1913 | "filterlim f sequentially F \<longleftrightarrow> filterlim (\<lambda>x. real (f x)) at_top F" (is "?lhs = ?rhs") | 
| 1914 | proof | |
| 1915 | assume ?lhs then show ?rhs | |
| 1916 | using filterlim_compose filterlim_real_sequentially by blast | |
| 1917 | next | |
| 1918 | assume R: ?rhs | |
| 1919 | show ?lhs | |
| 63556 | 1920 | proof - | 
| 1921 | have "filterlim (\<lambda>x. nat (floor (real (f x)))) sequentially F" | |
| 1922 | by (intro filterlim_compose[OF filterlim_nat_sequentially] | |
| 71720 | 1923 | filterlim_compose[OF filterlim_floor_sequentially] R) | 
| 63556 | 1924 | then show ?thesis by simp | 
| 1925 | qed | |
| 71720 | 1926 | qed | 
| 63556 | 1927 | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1928 | |
| 60758 | 1929 | subsubsection \<open>Limits of Sequences\<close> | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1930 | |
| 63545 | 1931 | lemma lim_sequentially: "X \<longlonglongrightarrow> L \<longleftrightarrow> (\<forall>r>0. \<exists>no. \<forall>n\<ge>no. dist (X n) L < r)" | 
| 1932 | for L :: "'a::metric_space" | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1933 | unfolding tendsto_iff eventually_sequentially .. | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1934 | |
| 60026 
41d81b4a0a21
Restored LIMSEQ_def as legacy binding. [The other changes are whitespace only.]
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 1935 | lemmas LIMSEQ_def = lim_sequentially (*legacy binding*) | 
| 
41d81b4a0a21
Restored LIMSEQ_def as legacy binding. [The other changes are whitespace only.]
 paulson <lp15@cam.ac.uk> parents: 
60017diff
changeset | 1936 | |
| 63545 | 1937 | lemma LIMSEQ_iff_nz: "X \<longlonglongrightarrow> L \<longleftrightarrow> (\<forall>r>0. \<exists>no>0. \<forall>n\<ge>no. dist (X n) L < r)" | 
| 1938 | for L :: "'a::metric_space" | |
| 60017 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 1939 | unfolding lim_sequentially by (metis Suc_leD zero_less_Suc) | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1940 | |
| 63545 | 1941 | lemma metric_LIMSEQ_I: "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. dist (X n) L < r) \<Longrightarrow> X \<longlonglongrightarrow> L" | 
| 1942 | for L :: "'a::metric_space" | |
| 1943 | by (simp add: lim_sequentially) | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1944 | |
| 63545 | 1945 | lemma metric_LIMSEQ_D: "X \<longlonglongrightarrow> L \<Longrightarrow> 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. dist (X n) L < r" | 
| 1946 | for L :: "'a::metric_space" | |
| 1947 | by (simp add: lim_sequentially) | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1948 | |
| 67673 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1949 | lemma LIMSEQ_norm_0: | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1950 | assumes "\<And>n::nat. norm (f n) < 1 / real (Suc n)" | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1951 | shows "f \<longlonglongrightarrow> 0" | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1952 | proof (rule metric_LIMSEQ_I) | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1953 | fix \<epsilon> :: "real" | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1954 | assume "\<epsilon> > 0" | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1955 | then obtain N::nat where "\<epsilon> > inverse N" "N > 0" | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1956 | by (metis neq0_conv real_arch_inverse) | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1957 | then have "norm (f n) < \<epsilon>" if "n \<ge> N" for n | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1958 | proof - | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1959 | have "1 / (Suc n) \<le> 1 / N" | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1960 | using \<open>0 < N\<close> inverse_of_nat_le le_SucI that by blast | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1961 | also have "\<dots> < \<epsilon>" | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1962 | by (metis (no_types) \<open>inverse (real N) < \<epsilon>\<close> inverse_eq_divide) | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1963 | finally show ?thesis | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1964 | by (meson assms less_eq_real_def not_le order_trans) | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1965 | qed | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1966 | then show "\<exists>no. \<forall>n\<ge>no. dist (f n) 0 < \<epsilon>" | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1967 | by auto | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1968 | qed | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1969 | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1970 | |
| 60758 | 1971 | subsubsection \<open>Limits of Functions\<close> | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1972 | |
| 63545 | 1973 | lemma LIM_def: "f \<midarrow>a\<rightarrow> L \<longleftrightarrow> (\<forall>r > 0. \<exists>s > 0. \<forall>x. x \<noteq> a \<and> dist x a < s \<longrightarrow> dist (f x) L < r)" | 
| 1974 | for a :: "'a::metric_space" and L :: "'b::metric_space" | |
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51531diff
changeset | 1975 | unfolding tendsto_iff eventually_at by simp | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1976 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1977 | lemma metric_LIM_I: | 
| 63545 | 1978 | "(\<And>r. 0 < r \<Longrightarrow> \<exists>s>0. \<forall>x. x \<noteq> a \<and> dist x a < s \<longrightarrow> dist (f x) L < r) \<Longrightarrow> f \<midarrow>a\<rightarrow> L" | 
| 1979 | for a :: "'a::metric_space" and L :: "'b::metric_space" | |
| 1980 | by (simp add: LIM_def) | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1981 | |
| 63545 | 1982 | lemma metric_LIM_D: "f \<midarrow>a\<rightarrow> L \<Longrightarrow> 0 < r \<Longrightarrow> \<exists>s>0. \<forall>x. x \<noteq> a \<and> dist x a < s \<longrightarrow> dist (f x) L < r" | 
| 1983 | for a :: "'a::metric_space" and L :: "'b::metric_space" | |
| 1984 | by (simp add: LIM_def) | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1985 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1986 | lemma metric_LIM_imp_LIM: | 
| 63545 | 1987 | fixes l :: "'a::metric_space" | 
| 1988 | and m :: "'b::metric_space" | |
| 1989 | assumes f: "f \<midarrow>a\<rightarrow> l" | |
| 1990 | and le: "\<And>x. x \<noteq> a \<Longrightarrow> dist (g x) m \<le> dist (f x) l" | |
| 1991 | shows "g \<midarrow>a\<rightarrow> m" | |
| 68594 | 1992 | by (rule metric_tendsto_imp_tendsto [OF f]) (auto simp: eventually_at_topological le) | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1993 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 1994 | lemma metric_LIM_equal2: | 
| 63545 | 1995 | fixes a :: "'a::metric_space" | 
| 68594 | 1996 | assumes "g \<midarrow>a\<rightarrow> l" "0 < R" | 
| 63545 | 1997 | and "\<And>x. x \<noteq> a \<Longrightarrow> dist x a < R \<Longrightarrow> f x = g x" | 
| 68594 | 1998 | shows "f \<midarrow>a\<rightarrow> l" | 
| 1999 | proof - | |
| 2000 | have "\<And>S. \<lbrakk>open S; l \<in> S; \<forall>\<^sub>F x in at a. g x \<in> S\<rbrakk> \<Longrightarrow> \<forall>\<^sub>F x in at a. f x \<in> S" | |
| 71720 | 2001 | apply (simp add: eventually_at) | 
| 2002 | by (metis assms(2) assms(3) dual_order.strict_trans linorder_neqE_linordered_idom) | |
| 68594 | 2003 | then show ?thesis | 
| 2004 | using assms by (simp add: tendsto_def) | |
| 2005 | qed | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2006 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2007 | lemma metric_LIM_compose2: | 
| 63545 | 2008 | fixes a :: "'a::metric_space" | 
| 2009 | assumes f: "f \<midarrow>a\<rightarrow> b" | |
| 2010 | and g: "g \<midarrow>b\<rightarrow> c" | |
| 2011 | and inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> f x \<noteq> b" | |
| 61976 | 2012 | shows "(\<lambda>x. g (f x)) \<midarrow>a\<rightarrow> c" | 
| 63545 | 2013 | using inj by (intro tendsto_compose_eventually[OF g f]) (auto simp: eventually_at) | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2014 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2015 | lemma metric_isCont_LIM_compose2: | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2016 | fixes f :: "'a :: metric_space \<Rightarrow> _" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2017 | assumes f [unfolded isCont_def]: "isCont f a" | 
| 63545 | 2018 | and g: "g \<midarrow>f a\<rightarrow> l" | 
| 2019 | and inj: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> f x \<noteq> f a" | |
| 61976 | 2020 | shows "(\<lambda>x. g (f x)) \<midarrow>a\<rightarrow> l" | 
| 63545 | 2021 | by (rule metric_LIM_compose2 [OF f g inj]) | 
| 2022 | ||
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2023 | |
| 60758 | 2024 | subsection \<open>Complete metric spaces\<close> | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2025 | |
| 60758 | 2026 | subsection \<open>Cauchy sequences\<close> | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2027 | |
| 62101 | 2028 | lemma (in metric_space) Cauchy_def: "Cauchy X = (\<forall>e>0. \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < e)" | 
| 2029 | proof - | |
| 63545 | 2030 |   have *: "eventually P (INF M. principal {(X m, X n) | n m. m \<ge> M \<and> n \<ge> M}) \<longleftrightarrow>
 | 
| 62101 | 2031 | (\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. P (X m, X n))" for P | 
| 63545 | 2032 | apply (subst eventually_INF_base) | 
| 2033 | subgoal by simp | |
| 2034 | subgoal for a b | |
| 62101 | 2035 | by (intro bexI[of _ "max a b"]) (auto simp: eventually_principal subset_eq) | 
| 63545 | 2036 | subgoal by (auto simp: eventually_principal, blast) | 
| 2037 | done | |
| 62101 | 2038 |   have "Cauchy X \<longleftrightarrow> (INF M. principal {(X m, X n) | n m. m \<ge> M \<and> n \<ge> M}) \<le> uniformity"
 | 
| 2039 | unfolding Cauchy_uniform_iff le_filter_def * .. | |
| 2040 | also have "\<dots> = (\<forall>e>0. \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < e)" | |
| 2041 | unfolding uniformity_dist le_INF_iff by (auto simp: * le_principal) | |
| 2042 | finally show ?thesis . | |
| 2043 | qed | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2044 | |
| 63545 | 2045 | lemma (in metric_space) Cauchy_altdef: "Cauchy f \<longleftrightarrow> (\<forall>e>0. \<exists>M. \<forall>m\<ge>M. \<forall>n>m. dist (f m) (f n) < e)" | 
| 2046 | (is "?lhs \<longleftrightarrow> ?rhs") | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2047 | proof | 
| 63545 | 2048 | assume ?rhs | 
| 2049 | show ?lhs | |
| 2050 | unfolding Cauchy_def | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2051 | proof (intro allI impI) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2052 | fix e :: real assume e: "e > 0" | 
| 63545 | 2053 | with \<open>?rhs\<close> obtain M where M: "m \<ge> M \<Longrightarrow> n > m \<Longrightarrow> dist (f m) (f n) < e" for m n | 
| 2054 | by blast | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2055 | have "dist (f m) (f n) < e" if "m \<ge> M" "n \<ge> M" for m n | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2056 | using M[of m n] M[of n m] e that by (cases m n rule: linorder_cases) (auto simp: dist_commute) | 
| 63545 | 2057 | then show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (f m) (f n) < e" | 
| 2058 | by blast | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2059 | qed | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2060 | next | 
| 63545 | 2061 | assume ?lhs | 
| 2062 | show ?rhs | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2063 | proof (intro allI impI) | 
| 63545 | 2064 | fix e :: real | 
| 2065 | assume e: "e > 0" | |
| 61799 | 2066 | with \<open>Cauchy f\<close> obtain M where "\<And>m n. m \<ge> M \<Longrightarrow> n \<ge> M \<Longrightarrow> dist (f m) (f n) < e" | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 2067 | unfolding Cauchy_def by blast | 
| 63545 | 2068 | then show "\<exists>M. \<forall>m\<ge>M. \<forall>n>m. dist (f m) (f n) < e" | 
| 2069 | by (intro exI[of _ M]) force | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2070 | qed | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2071 | qed | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2072 | |
| 66089 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2073 | lemma (in metric_space) Cauchy_altdef2: "Cauchy s \<longleftrightarrow> (\<forall>e>0. \<exists>N::nat. \<forall>n\<ge>N. dist(s n)(s N) < e)" (is "?lhs = ?rhs") | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2074 | proof | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2075 | assume "Cauchy s" | 
| 68594 | 2076 | then show ?rhs by (force simp: Cauchy_def) | 
| 66089 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2077 | next | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2078 | assume ?rhs | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2079 |     {
 | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2080 | fix e::real | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2081 | assume "e>0" | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2082 | with \<open>?rhs\<close> obtain N where N: "\<forall>n\<ge>N. dist (s n) (s N) < e/2" | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2083 | by (erule_tac x="e/2" in allE) auto | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2084 |       {
 | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2085 | fix n m | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2086 | assume nm: "N \<le> m \<and> N \<le> n" | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2087 | then have "dist (s m) (s n) < e" using N | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2088 | using dist_triangle_half_l[of "s m" "s N" "e" "s n"] | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2089 | by blast | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2090 | } | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2091 | then have "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (s m) (s n) < e" | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2092 | by blast | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2093 | } | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2094 | then have ?lhs | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2095 | unfolding Cauchy_def by blast | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2096 | then show ?lhs | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2097 | by blast | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2098 | qed | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 2099 | |
| 62101 | 2100 | lemma (in metric_space) metric_CauchyI: | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2101 | "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < e) \<Longrightarrow> Cauchy X" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2102 | by (simp add: Cauchy_def) | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2103 | |
| 63545 | 2104 | lemma (in metric_space) CauchyI': | 
| 2105 | "(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n>m. dist (X m) (X n) < e) \<Longrightarrow> Cauchy X" | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2106 | unfolding Cauchy_altdef by blast | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2107 | |
| 62101 | 2108 | lemma (in metric_space) metric_CauchyD: | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2109 | "Cauchy X \<Longrightarrow> 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < e" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2110 | by (simp add: Cauchy_def) | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2111 | |
| 62101 | 2112 | lemma (in metric_space) metric_Cauchy_iff2: | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2113 | "Cauchy X = (\<forall>j. (\<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. dist (X m) (X n) < inverse(real (Suc j))))" | 
| 68594 | 2114 | apply (auto simp add: Cauchy_def) | 
| 2115 | by (metis less_trans of_nat_Suc reals_Archimedean) | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2116 | |
| 63545 | 2117 | lemma Cauchy_iff2: "Cauchy X \<longleftrightarrow> (\<forall>j. (\<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. \<bar>X m - X n\<bar> < inverse (real (Suc j))))" | 
| 2118 | by (simp only: metric_Cauchy_iff2 dist_real_def) | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2119 | |
| 70723 
4e39d87c9737
imported new material mostly due to Sébastien Gouëzel
 paulson <lp15@cam.ac.uk> parents: 
70630diff
changeset | 2120 | lemma lim_1_over_n [tendsto_intros]: "((\<lambda>n. 1 / of_nat n) \<longlongrightarrow> (0::'a::real_normed_field)) sequentially" | 
| 62101 | 2121 | proof (subst lim_sequentially, intro allI impI exI) | 
| 70723 
4e39d87c9737
imported new material mostly due to Sébastien Gouëzel
 paulson <lp15@cam.ac.uk> parents: 
70630diff
changeset | 2122 | fix e::real and n | 
| 
4e39d87c9737
imported new material mostly due to Sébastien Gouëzel
 paulson <lp15@cam.ac.uk> parents: 
70630diff
changeset | 2123 | assume e: "e > 0" | 
| 62101 | 2124 | have "inverse e < of_nat (nat \<lceil>inverse e + 1\<rceil>)" by linarith | 
| 70723 
4e39d87c9737
imported new material mostly due to Sébastien Gouëzel
 paulson <lp15@cam.ac.uk> parents: 
70630diff
changeset | 2125 | also assume "n \<ge> nat \<lceil>inverse e + 1\<rceil>" | 
| 63545 | 2126 | finally show "dist (1 / of_nat n :: 'a) 0 < e" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 2127 | using e by (simp add: field_split_simps norm_divide) | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2128 | qed | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2129 | |
| 62101 | 2130 | lemma (in metric_space) complete_def: | 
| 2131 | shows "complete S = (\<forall>f. (\<forall>n. f n \<in> S) \<and> Cauchy f \<longrightarrow> (\<exists>l\<in>S. f \<longlonglongrightarrow> l))" | |
| 2132 | unfolding complete_uniform | |
| 2133 | proof safe | |
| 63545 | 2134 | fix f :: "nat \<Rightarrow> 'a" | 
| 2135 | assume f: "\<forall>n. f n \<in> S" "Cauchy f" | |
| 62101 | 2136 | and *: "\<forall>F\<le>principal S. F \<noteq> bot \<longrightarrow> cauchy_filter F \<longrightarrow> (\<exists>x\<in>S. F \<le> nhds x)" | 
| 2137 | then show "\<exists>l\<in>S. f \<longlonglongrightarrow> l" | |
| 2138 | unfolding filterlim_def using f | |
| 2139 | by (intro *[rule_format]) | |
| 2140 | (auto simp: filtermap_sequentually_ne_bot le_principal eventually_filtermap Cauchy_uniform) | |
| 2141 | next | |
| 63545 | 2142 | fix F :: "'a filter" | 
| 2143 | assume "F \<le> principal S" "F \<noteq> bot" "cauchy_filter F" | |
| 62101 | 2144 | assume seq: "\<forall>f. (\<forall>n. f n \<in> S) \<and> Cauchy f \<longrightarrow> (\<exists>l\<in>S. f \<longlonglongrightarrow> l)" | 
| 2145 | ||
| 63545 | 2146 | from \<open>F \<le> principal S\<close> \<open>cauchy_filter F\<close> | 
| 2147 | have FF_le: "F \<times>\<^sub>F F \<le> uniformity_on S" | |
| 62101 | 2148 | by (simp add: cauchy_filter_def principal_prod_principal[symmetric] prod_filter_mono) | 
| 2149 | ||
| 2150 | let ?P = "\<lambda>P e. eventually P F \<and> (\<forall>x. P x \<longrightarrow> x \<in> S) \<and> (\<forall>x y. P x \<longrightarrow> P y \<longrightarrow> dist x y < e)" | |
| 63545 | 2151 | have P: "\<exists>P. ?P P \<epsilon>" if "0 < \<epsilon>" for \<epsilon> :: real | 
| 2152 | proof - | |
| 2153 | from that have "eventually (\<lambda>(x, y). x \<in> S \<and> y \<in> S \<and> dist x y < \<epsilon>) (uniformity_on S)" | |
| 2154 | by (auto simp: eventually_inf_principal eventually_uniformity_metric) | |
| 2155 | from filter_leD[OF FF_le this] show ?thesis | |
| 2156 | by (auto simp: eventually_prod_same) | |
| 2157 | qed | |
| 62101 | 2158 | |
| 2159 | have "\<exists>P. \<forall>n. ?P (P n) (1 / Suc n) \<and> P (Suc n) \<le> P n" | |
| 2160 | proof (rule dependent_nat_choice) | |
| 2161 | show "\<exists>P. ?P P (1 / Suc 0)" | |
| 2162 | using P[of 1] by auto | |
| 2163 | next | |
| 2164 | fix P n assume "?P P (1/Suc n)" | |
| 2165 | moreover obtain Q where "?P Q (1 / Suc (Suc n))" | |
| 2166 | using P[of "1/Suc (Suc n)"] by auto | |
| 2167 | ultimately show "\<exists>Q. ?P Q (1 / Suc (Suc n)) \<and> Q \<le> P" | |
| 2168 | by (intro exI[of _ "\<lambda>x. P x \<and> Q x"]) (auto simp: eventually_conj_iff) | |
| 2169 | qed | |
| 63545 | 2170 | then obtain P where P: "eventually (P n) F" "P n x \<Longrightarrow> x \<in> S" | 
| 2171 | "P n x \<Longrightarrow> P n y \<Longrightarrow> dist x y < 1 / Suc n" "P (Suc n) \<le> P n" | |
| 2172 | for n x y | |
| 62101 | 2173 | by metis | 
| 2174 | have "antimono P" | |
| 76055 
8d56461f85ec
moved antimono to Fun and redefined it as an abbreviation
 desharna parents: 
74475diff
changeset | 2175 | using P(4) by (rule decseq_SucI) | 
| 62101 | 2176 | |
| 63545 | 2177 | obtain X where X: "P n (X n)" for n | 
| 62101 | 2178 | using P(1)[THEN eventually_happens'[OF \<open>F \<noteq> bot\<close>]] by metis | 
| 2179 | have "Cauchy X" | |
| 2180 | unfolding metric_Cauchy_iff2 inverse_eq_divide | |
| 2181 | proof (intro exI allI impI) | |
| 63545 | 2182 | fix j m n :: nat | 
| 2183 | assume "j \<le> m" "j \<le> n" | |
| 62101 | 2184 | with \<open>antimono P\<close> X have "P j (X m)" "P j (X n)" | 
| 2185 | by (auto simp: antimono_def) | |
| 2186 | then show "dist (X m) (X n) < 1 / Suc j" | |
| 2187 | by (rule P) | |
| 2188 | qed | |
| 2189 | moreover have "\<forall>n. X n \<in> S" | |
| 2190 | using P(2) X by auto | |
| 2191 | ultimately obtain x where "X \<longlonglongrightarrow> x" "x \<in> S" | |
| 2192 | using seq by blast | |
| 2193 | ||
| 2194 | show "\<exists>x\<in>S. F \<le> nhds x" | |
| 2195 | proof (rule bexI) | |
| 63545 | 2196 | have "eventually (\<lambda>y. dist y x < e) F" if "0 < e" for e :: real | 
| 2197 | proof - | |
| 2198 | from that have "(\<lambda>n. 1 / Suc n :: real) \<longlonglongrightarrow> 0 \<and> 0 < e / 2" | |
| 71827 | 2199 | by (subst filterlim_sequentially_Suc) (auto intro!: lim_1_over_n) | 
| 62101 | 2200 | then have "\<forall>\<^sub>F n in sequentially. dist (X n) x < e / 2 \<and> 1 / Suc n < e / 2" | 
| 63545 | 2201 | using \<open>X \<longlonglongrightarrow> x\<close> | 
| 2202 | unfolding tendsto_iff order_tendsto_iff[where 'a=real] eventually_conj_iff | |
| 2203 | by blast | |
| 62101 | 2204 | then obtain n where "dist x (X n) < e / 2" "1 / Suc n < e / 2" | 
| 2205 | by (auto simp: eventually_sequentially dist_commute) | |
| 63545 | 2206 | show ?thesis | 
| 62101 | 2207 | using \<open>eventually (P n) F\<close> | 
| 2208 | proof eventually_elim | |
| 63545 | 2209 | case (elim y) | 
| 62101 | 2210 | then have "dist y (X n) < 1 / Suc n" | 
| 2211 | by (intro X P) | |
| 2212 | also have "\<dots> < e / 2" by fact | |
| 2213 | finally show "dist y x < e" | |
| 2214 | by (rule dist_triangle_half_l) fact | |
| 63545 | 2215 | qed | 
| 2216 | qed | |
| 62101 | 2217 | then show "F \<le> nhds x" | 
| 2218 | unfolding nhds_metric le_INF_iff le_principal by auto | |
| 2219 | qed fact | |
| 2220 | qed | |
| 2221 | ||
| 68594 | 2222 | text\<open>apparently unused\<close> | 
| 62101 | 2223 | lemma (in metric_space) totally_bounded_metric: | 
| 2224 |   "totally_bounded S \<longleftrightarrow> (\<forall>e>0. \<exists>k. finite k \<and> S \<subseteq> (\<Union>x\<in>k. {y. dist x y < e}))"
 | |
| 68594 | 2225 | unfolding totally_bounded_def eventually_uniformity_metric imp_ex | 
| 62101 | 2226 | apply (subst all_comm) | 
| 68594 | 2227 | apply (intro arg_cong[where f=All] ext, safe) | 
| 62101 | 2228 | subgoal for e | 
| 2229 | apply (erule allE[of _ "\<lambda>(x, y). dist x y < e"]) | |
| 2230 | apply auto | |
| 2231 | done | |
| 2232 | subgoal for e P k | |
| 2233 | apply (intro exI[of _ k]) | |
| 2234 | apply (force simp: subset_eq) | |
| 2235 | done | |
| 2236 | done | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2237 | |
| 63545 | 2238 | |
| 74475 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2239 | setup \<open>Sign.add_const_constraint (\<^const_name>\<open>dist\<close>, SOME \<^typ>\<open>'a::dist \<Rightarrow> 'a \<Rightarrow> real\<close>)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2240 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2241 | (* Contributed by Dominique Unruh *) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2242 | lemma cauchy_filter_metric: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2243 |   fixes F :: "'a::{uniformity_dist,uniform_space} filter"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2244 | shows "cauchy_filter F \<longleftrightarrow> (\<forall>e. e>0 \<longrightarrow> (\<exists>P. eventually P F \<and> (\<forall>x y. P x \<and> P y \<longrightarrow> dist x y < e)))" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2245 | proof (unfold cauchy_filter_def le_filter_def, auto) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2246 | assume assm: \<open>\<forall>e>0. \<exists>P. eventually P F \<and> (\<forall>x y. P x \<and> P y \<longrightarrow> dist x y < e)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2247 | then show \<open>eventually P uniformity \<Longrightarrow> eventually P (F \<times>\<^sub>F F)\<close> for P | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2248 | apply (auto simp: eventually_uniformity_metric) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2249 | using eventually_prod_same by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2250 | next | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2251 | fix e :: real | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2252 | assume \<open>e > 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2253 | assume asm: \<open>\<forall>P. eventually P uniformity \<longrightarrow> eventually P (F \<times>\<^sub>F F)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2254 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2255 | define P where \<open>P \<equiv> \<lambda>(x,y :: 'a). dist x y < e\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2256 | with asm \<open>e > 0\<close> have \<open>eventually P (F \<times>\<^sub>F F)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2257 | by (metis case_prod_conv eventually_uniformity_metric) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2258 | then | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2259 | show \<open>\<exists>P. eventually P F \<and> (\<forall>x y. P x \<and> P y \<longrightarrow> dist x y < e)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2260 | by (auto simp add: eventually_prod_same P_def) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2261 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2262 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2263 | (* Contributed by Dominique Unruh *) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2264 | lemma cauchy_filter_metric_filtermap: | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2265 |   fixes f :: "'a \<Rightarrow> 'b::{uniformity_dist,uniform_space}"
 | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2266 | shows "cauchy_filter (filtermap f F) \<longleftrightarrow> (\<forall>e. e>0 \<longrightarrow> (\<exists>P. eventually P F \<and> (\<forall>x y. P x \<and> P y \<longrightarrow> dist (f x) (f y) < e)))" | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2267 | proof (subst cauchy_filter_metric, intro iffI allI impI) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2268 | assume \<open>\<forall>e>0. \<exists>P. eventually P (filtermap f F) \<and> (\<forall>x y. P x \<and> P y \<longrightarrow> dist x y < e)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2269 | then show \<open>e>0 \<Longrightarrow> \<exists>P. eventually P F \<and> (\<forall>x y. P x \<and> P y \<longrightarrow> dist (f x) (f y) < e)\<close> for e | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2270 | unfolding eventually_filtermap by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2271 | next | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2272 | assume asm: \<open>\<forall>e>0. \<exists>P. eventually P F \<and> (\<forall>x y. P x \<and> P y \<longrightarrow> dist (f x) (f y) < e)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2273 | fix e::real assume \<open>e > 0\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2274 | then obtain P where \<open>eventually P F\<close> and PPe: \<open>P x \<and> P y \<longrightarrow> dist (f x) (f y) < e\<close> for x y | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2275 | using asm by blast | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2276 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2277 | show \<open>\<exists>P. eventually P (filtermap f F) \<and> (\<forall>x y. P x \<and> P y \<longrightarrow> dist x y < e)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2278 | apply (rule exI[of _ \<open>\<lambda>x. \<exists>y. P y \<and> x = f y\<close>]) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2279 | using PPe \<open>eventually P F\<close> apply (auto simp: eventually_filtermap) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2280 | by (smt (verit, ccfv_SIG) eventually_elim2) | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2281 | qed | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2282 | |
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2283 | setup \<open>Sign.add_const_constraint (\<^const_name>\<open>dist\<close>, SOME \<^typ>\<open>'a::metric_space \<Rightarrow> 'a \<Rightarrow> real\<close>)\<close> | 
| 
409ca22dee4c
new notion of infinite sums in HOL-Analysis, ordering on complex numbers
 eberlm <eberlm@in.tum.de> parents: 
74007diff
changeset | 2284 | |
| 60758 | 2285 | subsubsection \<open>Cauchy Sequences are Convergent\<close> | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2286 | |
| 62101 | 2287 | (* TODO: update to uniform_space *) | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2288 | class complete_space = metric_space + | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2289 | assumes Cauchy_convergent: "Cauchy X \<Longrightarrow> convergent X" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2290 | |
| 63545 | 2291 | lemma Cauchy_convergent_iff: "Cauchy X \<longleftrightarrow> convergent X" | 
| 2292 | for X :: "nat \<Rightarrow> 'a::complete_space" | |
| 2293 | by (blast intro: Cauchy_convergent convergent_Cauchy) | |
| 2294 | ||
| 67727 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 2295 | text \<open>To prove that a Cauchy sequence converges, it suffices to show that a subsequence converges.\<close> | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 2296 | |
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 2297 | lemma Cauchy_converges_subseq: | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 2298 | fixes u::"nat \<Rightarrow> 'a::metric_space" | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 2299 | assumes "Cauchy u" | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 2300 | "strict_mono r" | 
| 68594 | 2301 | "(u \<circ> r) \<longlonglongrightarrow> l" | 
| 67727 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 2302 | shows "u \<longlonglongrightarrow> l" | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 2303 | proof - | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 2304 | have *: "eventually (\<lambda>n. dist (u n) l < e) sequentially" if "e > 0" for e | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 2305 | proof - | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 2306 | have "e/2 > 0" using that by auto | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 2307 | then obtain N1 where N1: "\<And>m n. m \<ge> N1 \<Longrightarrow> n \<ge> N1 \<Longrightarrow> dist (u m) (u n) < e/2" | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 2308 | using \<open>Cauchy u\<close> unfolding Cauchy_def by blast | 
| 68594 | 2309 | obtain N2 where N2: "\<And>n. n \<ge> N2 \<Longrightarrow> dist ((u \<circ> r) n) l < e / 2" | 
| 2310 | using order_tendstoD(2)[OF iffD1[OF tendsto_dist_iff \<open>(u \<circ> r) \<longlonglongrightarrow> l\<close>] \<open>e/2 > 0\<close>] | |
| 67727 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 2311 | unfolding eventually_sequentially by auto | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 2312 | have "dist (u n) l < e" if "n \<ge> max N1 N2" for n | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 2313 | proof - | 
| 68594 | 2314 | have "dist (u n) l \<le> dist (u n) ((u \<circ> r) n) + dist ((u \<circ> r) n) l" | 
| 67727 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 2315 | by (rule dist_triangle) | 
| 68594 | 2316 | also have "\<dots> < e/2 + e/2" | 
| 71720 | 2317 | proof (intro add_strict_mono) | 
| 2318 | show "dist (u n) ((u \<circ> r) n) < e / 2" | |
| 2319 | using N1[of n "r n"] N2[of n] that unfolding comp_def | |
| 2320 | by (meson assms(2) le_trans max.bounded_iff strict_mono_imp_increasing) | |
| 2321 | show "dist ((u \<circ> r) n) l < e / 2" | |
| 2322 | using N2 that by auto | |
| 2323 | qed | |
| 67727 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 2324 | finally show ?thesis by simp | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 2325 | qed | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 2326 | then show ?thesis unfolding eventually_sequentially by blast | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 2327 | qed | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 2328 | have "(\<lambda>n. dist (u n) l) \<longlonglongrightarrow> 0" | 
| 71720 | 2329 | by (simp add: less_le_trans * order_tendstoI) | 
| 67727 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 2330 | then show ?thesis using tendsto_dist_iff by auto | 
| 
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
 immler parents: 
67706diff
changeset | 2331 | qed | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2332 | |
| 60758 | 2333 | subsection \<open>The set of real numbers is a complete metric space\<close> | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2334 | |
| 60758 | 2335 | text \<open> | 
| 63545 | 2336 | Proof that Cauchy sequences converge based on the one from | 
| 63680 | 2337 | \<^url>\<open>http://pirate.shu.edu/~wachsmut/ira/numseq/proofs/cauconv.html\<close> | 
| 60758 | 2338 | \<close> | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2339 | |
| 60758 | 2340 | text \<open> | 
| 69593 | 2341 | If sequence \<^term>\<open>X\<close> is Cauchy, then its limit is the lub of | 
| 2342 |   \<^term>\<open>{r::real. \<exists>N. \<forall>n\<ge>N. r < X n}\<close>
 | |
| 60758 | 2343 | \<close> | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2344 | lemma increasing_LIMSEQ: | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2345 | fixes f :: "nat \<Rightarrow> real" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2346 | assumes inc: "\<And>n. f n \<le> f (Suc n)" | 
| 63545 | 2347 | and bdd: "\<And>n. f n \<le> l" | 
| 2348 | and en: "\<And>e. 0 < e \<Longrightarrow> \<exists>n. l \<le> f n + e" | |
| 61969 | 2349 | shows "f \<longlonglongrightarrow> l" | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2350 | proof (rule increasing_tendsto) | 
| 63545 | 2351 | fix x | 
| 2352 | assume "x < l" | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2353 | with dense[of 0 "l - x"] obtain e where "0 < e" "e < l - x" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2354 | by auto | 
| 60758 | 2355 | from en[OF \<open>0 < e\<close>] obtain n where "l - e \<le> f n" | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2356 | by (auto simp: field_simps) | 
| 63545 | 2357 | with \<open>e < l - x\<close> \<open>0 < e\<close> have "x < f n" | 
| 2358 | by simp | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2359 | with incseq_SucI[of f, OF inc] show "eventually (\<lambda>n. x < f n) sequentially" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2360 | by (auto simp: eventually_sequentially incseq_def intro: less_le_trans) | 
| 63545 | 2361 | qed (use bdd in auto) | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2362 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2363 | lemma real_Cauchy_convergent: | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2364 | fixes X :: "nat \<Rightarrow> real" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2365 | assumes X: "Cauchy X" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2366 | shows "convergent X" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2367 | proof - | 
| 63040 | 2368 |   define S :: "real set" where "S = {x. \<exists>N. \<forall>n\<ge>N. x < X n}"
 | 
| 63545 | 2369 | then have mem_S: "\<And>N x. \<forall>n\<ge>N. x < X n \<Longrightarrow> x \<in> S" | 
| 2370 | by auto | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2371 | |
| 63545 | 2372 | have bound_isUb: "y \<le> x" if N: "\<forall>n\<ge>N. X n < x" and "y \<in> S" for N and x y :: real | 
| 2373 | proof - | |
| 2374 | from that have "\<exists>M. \<forall>n\<ge>M. y < X n" | |
| 2375 | by (simp add: S_def) | |
| 2376 | then obtain M where "\<forall>n\<ge>M. y < X n" .. | |
| 2377 | then have "y < X (max M N)" by simp | |
| 2378 | also have "\<dots> < x" using N by simp | |
| 2379 | finally show ?thesis by (rule order_less_imp_le) | |
| 2380 | qed | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2381 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2382 | obtain N where "\<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m) (X n) < 1" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2383 | using X[THEN metric_CauchyD, OF zero_less_one] by auto | 
| 63545 | 2384 | then have N: "\<forall>n\<ge>N. dist (X n) (X N) < 1" by simp | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54230diff
changeset | 2385 |   have [simp]: "S \<noteq> {}"
 | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54230diff
changeset | 2386 | proof (intro exI ex_in_conv[THEN iffD1]) | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2387 | from N have "\<forall>n\<ge>N. X N - 1 < X n" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2388 | by (simp add: abs_diff_less_iff dist_real_def) | 
| 63545 | 2389 | then show "X N - 1 \<in> S" by (rule mem_S) | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2390 | qed | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54230diff
changeset | 2391 | have [simp]: "bdd_above S" | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2392 | proof | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2393 | from N have "\<forall>n\<ge>N. X n < X N + 1" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2394 | by (simp add: abs_diff_less_iff dist_real_def) | 
| 63545 | 2395 | then show "\<And>s. s \<in> S \<Longrightarrow> s \<le> X N + 1" | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2396 | by (rule bound_isUb) | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2397 | qed | 
| 61969 | 2398 | have "X \<longlonglongrightarrow> Sup S" | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2399 | proof (rule metric_LIMSEQ_I) | 
| 63545 | 2400 | fix r :: real | 
| 2401 | assume "0 < r" | |
| 2402 | then have r: "0 < r/2" by simp | |
| 2403 | obtain N where "\<forall>n\<ge>N. \<forall>m\<ge>N. dist (X n) (X m) < r/2" | |
| 2404 | using metric_CauchyD [OF X r] by auto | |
| 2405 | then have "\<forall>n\<ge>N. dist (X n) (X N) < r/2" by simp | |
| 2406 | then have N: "\<forall>n\<ge>N. X N - r/2 < X n \<and> X n < X N + r/2" | |
| 2407 | by (simp only: dist_real_def abs_diff_less_iff) | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2408 | |
| 63545 | 2409 | from N have "\<forall>n\<ge>N. X N - r/2 < X n" by blast | 
| 2410 | then have "X N - r/2 \<in> S" by (rule mem_S) | |
| 2411 | then have 1: "X N - r/2 \<le> Sup S" by (simp add: cSup_upper) | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2412 | |
| 63545 | 2413 | from N have "\<forall>n\<ge>N. X n < X N + r/2" by blast | 
| 2414 | from bound_isUb[OF this] | |
| 2415 | have 2: "Sup S \<le> X N + r/2" | |
| 2416 | by (intro cSup_least) simp_all | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2417 | |
| 63545 | 2418 | show "\<exists>N. \<forall>n\<ge>N. dist (X n) (Sup S) < r" | 
| 2419 | proof (intro exI allI impI) | |
| 2420 | fix n | |
| 2421 | assume n: "N \<le> n" | |
| 2422 | from N n have "X n < X N + r/2" and "X N - r/2 < X n" | |
| 2423 | by simp_all | |
| 2424 | then show "dist (X n) (Sup S) < r" using 1 2 | |
| 2425 | by (simp add: abs_diff_less_iff dist_real_def) | |
| 2426 | qed | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2427 | qed | 
| 63545 | 2428 | then show ?thesis by (auto simp: convergent_def) | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2429 | qed | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2430 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2431 | instance real :: complete_space | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2432 | by intro_classes (rule real_Cauchy_convergent) | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2433 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2434 | class banach = real_normed_vector + complete_space | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2435 | |
| 61169 | 2436 | instance real :: banach .. | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2437 | |
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2438 | lemma tendsto_at_topI_sequentially: | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 2439 | fixes f :: "real \<Rightarrow> 'b::first_countable_topology" | 
| 61969 | 2440 | assumes *: "\<And>X. filterlim X at_top sequentially \<Longrightarrow> (\<lambda>n. f (X n)) \<longlonglongrightarrow> y" | 
| 61973 | 2441 | shows "(f \<longlongrightarrow> y) at_top" | 
| 57448 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 2442 | proof - | 
| 63545 | 2443 | obtain A where A: "decseq A" "open (A n)" "y \<in> A n" "nhds y = (INF n. principal (A n))" for n | 
| 2444 | by (rule nhds_countable[of y]) (rule that) | |
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 2445 | |
| 57448 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 2446 | have "\<forall>m. \<exists>k. \<forall>x\<ge>k. f x \<in> A m" | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 2447 | proof (rule ccontr) | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 2448 | assume "\<not> (\<forall>m. \<exists>k. \<forall>x\<ge>k. f x \<in> A m)" | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 2449 | then obtain m where "\<And>k. \<exists>x\<ge>k. f x \<notin> A m" | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 2450 | by auto | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 2451 | then have "\<exists>X. \<forall>n. (f (X n) \<notin> A m) \<and> max n (X n) + 1 \<le> X (Suc n)" | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 2452 | by (intro dependent_nat_choice) (auto simp del: max.bounded_iff) | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 2453 | then obtain X where X: "\<And>n. f (X n) \<notin> A m" "\<And>n. max n (X n) + 1 \<le> X (Suc n)" | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 2454 | by auto | 
| 63545 | 2455 | have "1 \<le> n \<Longrightarrow> real n \<le> X n" for n | 
| 2456 | using X[of "n - 1"] by auto | |
| 57448 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 2457 | then have "filterlim X at_top sequentially" | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 2458 | by (force intro!: filterlim_at_top_mono[OF filterlim_real_sequentially] | 
| 63545 | 2459 | simp: eventually_sequentially) | 
| 57448 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 2460 | from topological_tendstoD[OF *[OF this] A(2, 3), of m] X(1) show False | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 2461 | by auto | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 2462 | qed | 
| 63545 | 2463 | then obtain k where "k m \<le> x \<Longrightarrow> f x \<in> A m" for m x | 
| 57448 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 2464 | by metis | 
| 
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
 hoelzl parents: 
57418diff
changeset | 2465 | then show ?thesis | 
| 63545 | 2466 | unfolding at_top_def A by (intro filterlim_base[where i=k]) auto | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 2467 | qed | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 2468 | |
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56889diff
changeset | 2469 | lemma tendsto_at_topI_sequentially_real: | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2470 | fixes f :: "real \<Rightarrow> real" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2471 | assumes mono: "mono f" | 
| 63545 | 2472 | and limseq: "(\<lambda>n. f (real n)) \<longlonglongrightarrow> y" | 
| 61973 | 2473 | shows "(f \<longlongrightarrow> y) at_top" | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2474 | proof (rule tendstoI) | 
| 63545 | 2475 | fix e :: real | 
| 2476 | assume "0 < e" | |
| 2477 | with limseq obtain N :: nat where N: "N \<le> n \<Longrightarrow> \<bar>f (real n) - y\<bar> < e" for n | |
| 60017 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 2478 | by (auto simp: lim_sequentially dist_real_def) | 
| 63545 | 2479 | have le: "f x \<le> y" for x :: real | 
| 2480 | proof - | |
| 53381 | 2481 | obtain n where "x \<le> real_of_nat n" | 
| 62623 
dbc62f86a1a9
rationalisation of theorem names esp about "real Archimedian" etc.
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 2482 | using real_arch_simple[of x] .. | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2483 | note monoD[OF mono this] | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2484 | also have "f (real_of_nat n) \<le> y" | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61609diff
changeset | 2485 | by (rule LIMSEQ_le_const[OF limseq]) (auto intro!: exI[of _ n] monoD[OF mono]) | 
| 63545 | 2486 | finally show ?thesis . | 
| 2487 | qed | |
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2488 | have "eventually (\<lambda>x. real N \<le> x) at_top" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2489 | by (rule eventually_ge_at_top) | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2490 | then show "eventually (\<lambda>x. dist (f x) y < e) at_top" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2491 | proof eventually_elim | 
| 63545 | 2492 | case (elim x) | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2493 | with N[of N] le have "y - f (real N) < e" by auto | 
| 63545 | 2494 | moreover note monoD[OF mono elim] | 
| 51531 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2495 | ultimately show "dist (f x) y < e" | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2496 | using le[of x] by (auto simp: dist_real_def field_simps) | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2497 | qed | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2498 | qed | 
| 
f415febf4234
remove Metric_Spaces and move its content into Limits and Real_Vector_Spaces
 hoelzl parents: 
51524diff
changeset | 2499 | |
| 20504 
6342e872e71d
formalization of vector spaces and algebras over the real numbers
 huffman parents: diff
changeset | 2500 | end |