diff --git a/thys/Affine_Arithmetic/Affine_Arithmetic_Auxiliarities.thy b/thys/Affine_Arithmetic/Affine_Arithmetic_Auxiliarities.thy --- a/thys/Affine_Arithmetic/Affine_Arithmetic_Auxiliarities.thy +++ b/thys/Affine_Arithmetic/Affine_Arithmetic_Auxiliarities.thy @@ -1,91 +1,91 @@ theory Affine_Arithmetic_Auxiliarities -imports "HOL-Analysis.Ordered_Euclidean_Space" "HOL-Analysis.Determinants" +imports "HOL-Analysis.Multivariate_Analysis" begin subsection \@{term sum_list}\ lemma sum_list_nth_eqI: fixes xs ys::"'a::monoid_add list" shows "length xs = length ys \ (\x y. (x, y) \ set (zip xs ys) \ x = y) \ sum_list xs = sum_list ys" by (induct xs ys rule: list_induct2) auto lemma fst_sum_list: "fst (sum_list xs) = sum_list (map fst xs)" by (induct xs) auto lemma snd_sum_list: "snd (sum_list xs) = sum_list (map snd xs)" by (induct xs) auto lemma take_greater_eqI: "take c xs = take c ys \ c \ a \ take a xs = take a ys" proof (induct xs arbitrary: a c ys) case (Cons x xs) note ICons = Cons thus ?case proof (cases a) case (Suc b) thus ?thesis using Cons(2,3) proof (cases ys) case (Cons z zs) from ICons obtain d where c: "c = Suc d" by (auto simp: Cons Suc dest!: Suc_le_D) show ?thesis using ICons(2,3) by (auto simp: Suc Cons c intro: ICons(1)) qed simp qed simp qed (metis le_0_eq take_eq_Nil) lemma take_max_eqD: "take (max a b) xs = take (max a b) ys \ take a xs = take a ys \ take b xs = take b ys" by (metis max.cobounded1 max.cobounded2 take_greater_eqI) lemma take_Suc_eq: "take (Suc n) xs = (if n < length xs then take n xs @ [xs ! n] else xs)" by (auto simp: take_Suc_conv_app_nth) subsection \Radiant and Degree\ definition "rad_of w = w * pi / 180" definition "deg_of w = 180 * w / pi" lemma rad_of_inverse[simp]: "deg_of (rad_of w) = w" and deg_of_inverse[simp]: "rad_of (deg_of w) = w" by (auto simp: deg_of_def rad_of_def) lemma deg_of_monoI: "x \ y \ deg_of x \ deg_of y" by (auto simp: deg_of_def intro!: divide_right_mono) lemma rad_of_monoI: "x \ y \ rad_of x \ rad_of y" by (auto simp: rad_of_def) lemma deg_of_strict_monoI: "x < y \ deg_of x < deg_of y" by (auto simp: deg_of_def intro!: divide_strict_right_mono) lemma rad_of_strict_monoI: "x < y \ rad_of x < rad_of y" by (auto simp: rad_of_def) lemma deg_of_mono[simp]: "deg_of x \ deg_of y \ x \ y" using rad_of_monoI by (fastforce intro!: deg_of_monoI) lemma rad_of_mono[simp]: "rad_of x \ rad_of y \ x \ y" using rad_of_monoI by (fastforce intro!: deg_of_monoI) lemma deg_of_strict_mono[simp]: "deg_of x < deg_of y \ x < y" using rad_of_strict_monoI by (fastforce intro!: deg_of_strict_monoI) lemma rad_of_strict_mono[simp]: "rad_of x < rad_of y \ x < y" using rad_of_strict_monoI by (fastforce intro!: deg_of_strict_monoI) lemma rad_of_lt_iff: "rad_of d < r \ d < deg_of r" and rad_of_gt_iff: "rad_of d > r \ d > deg_of r" and rad_of_le_iff: "rad_of d \ r \ d \ deg_of r" and rad_of_ge_iff: "rad_of d \ r \ d \ deg_of r" using rad_of_strict_mono[of d "deg_of r"] rad_of_mono[of d "deg_of r"] by auto end diff --git a/thys/Affine_Arithmetic/Affine_Form.thy b/thys/Affine_Arithmetic/Affine_Form.thy --- a/thys/Affine_Arithmetic/Affine_Form.thy +++ b/thys/Affine_Arithmetic/Affine_Form.thy @@ -1,1989 +1,1989 @@ section \Affine Form\ theory Affine_Form imports - "HOL-Analysis.Ordered_Euclidean_Space" + "HOL-Analysis.Multivariate_Analysis" "HOL-Library.Permutation" Affine_Arithmetic_Auxiliarities Executable_Euclidean_Space begin subsection \Auxiliary developments\ lemma sum_list_mono: fixes xs ys::"'a::ordered_ab_group_add list" shows "length xs = length ys \ (\x y. (x, y) \ set (zip xs ys) \ x \ y) \ sum_list xs \ sum_list ys" by (induct xs ys rule: list_induct2) (auto simp: algebra_simps intro: add_mono) lemma fixes xs::"'a::ordered_comm_monoid_add list" shows sum_list_nonneg: "(\x. x \ set xs \ x \ 0) \ sum_list xs \ 0" by (induct xs) (auto intro!: add_nonneg_nonneg) lemma map_filter: "map f (filter (\x. P (f x)) xs) = filter P (map f xs)" by (induct xs) simp_all lemma map_of_zip_upto2_length_eq_nth: assumes "distinct B" assumes "i < length B" shows "(map_of (zip B [0.. (i, a) \ set xs \ (i, b) \ set xs \ a = b" by (metis (lifting) map_of_is_SomeI option.inject) lemma length_filter_snd_zip: "length ys = length xs \ length (filter (p \ snd) (zip ys xs)) = length (filter p xs)" by (induct ys xs rule: list_induct2) (auto ) lemma filter_snd_nth: "length ys = length xs \ n < length (filter p xs) \ snd (filter (p \ snd) (zip ys xs) ! n) = filter p xs ! n" by (induct ys xs arbitrary: n rule: list_induct2) (auto simp: o_def nth_Cons split: nat.split) lemma distinct_map_snd_fst_eqD: "distinct (map snd xs) \ (i, a) \ set xs \ (j, a) \ set xs \ i = j" by (metis Pair_inject inj_on_contraD snd_conv distinct_map) lemma map_of_mapk_inj_on_SomeI: "inj_on f (fst ` (set t)) \ map_of t k = Some x \ map_of (map (case_prod (\k. Pair (f k))) t) (f k) = Some x" by (induct t) (auto simp add: inj_on_def dest!: map_of_SomeD split: if_split_asm) lemma map_abs_nonneg[simp]: fixes xs::"'a::ordered_ab_group_add_abs list" shows "list_all (\x. x \ 0) xs \ map abs xs = xs" by (induct xs) auto lemma the_inv_into_image_eq: "inj_on f A \ Y \ f ` A \ the_inv_into A f ` Y = f -` Y \ A" using f_the_inv_into_f the_inv_into_f_f[where f = f and A = A] by force lemma image_fst_zip: "length ys = length xs \ fst ` set (zip ys xs) = set ys" by (metis dom_map_of_conv_image_fst dom_map_of_zip) lemma inj_on_fst_set_zip_distinct[simp]: "distinct xs \ length xs = length ys \ inj_on fst (set (zip xs ys))" by (force simp add: in_set_zip distinct_conv_nth intro!: inj_onI) lemma mem_greaterThanLessThan_absI: fixes x::real assumes "abs x < 1" shows "x \ {-1 <..< 1}" using assms by (auto simp: abs_real_def split: if_split_asm) lemma minus_one_less_divideI: "b > 0 \ -b < a \ -1 < a / (b::real)" by (auto simp: field_simps) lemma divide_less_oneI: "b > 0 \ b > a \ a / (b::real) < 1" by (auto simp: field_simps) lemma closed_segment_real: fixes a b::real shows "closed_segment a b = (if a \ b then {a .. b} else {b .. a})" (is "_ = ?if") proof safe fix x assume "x \ closed_segment a b" from segment_bound[OF this] show "x \ ?if" by (auto simp: abs_real_def split: if_split_asm) next fix x assume "x \ ?if" thus "x \ closed_segment a b" by (auto simp: closed_segment_def intro!: exI[where x="(x - a)/(b - a)"] simp: divide_simps algebra_simps) qed subsection \Partial Deviations\ typedef (overloaded) 'a pdevs = "{x::nat \ 'a::zero. finite {i. x i \ 0}}" \ \TODO: unify with polynomials\ morphisms pdevs_apply Abs_pdev by (auto intro!: exI[where x="\x. 0"]) setup_lifting type_definition_pdevs lemma pdevs_eqI: "(\i. pdevs_apply x i = pdevs_apply y i) \ x = y" by transfer auto definition pdevs_val :: "(nat \ real) \ 'a::real_normed_vector pdevs \ 'a" where "pdevs_val e x = (\i. e i *\<^sub>R pdevs_apply x i)" definition valuate:: "((nat \ real) \ 'a) \ 'a set" where "valuate x = x ` (UNIV \ {-1 .. 1})" lemma valuate_ex: "x \ valuate f \ (\e. (\i. e i \ {-1 .. 1}) \ x = f e)" unfolding valuate_def by (auto simp add: valuate_def Pi_iff) blast instantiation pdevs :: (equal) equal begin definition equal_pdevs::"'a pdevs \ 'a pdevs \ bool" where "equal_pdevs a b \ a = b" instance proof qed (simp add: equal_pdevs_def) end subsection \Affine Forms\ text \The data structure of affine forms represents particular sets, zonotopes\ type_synonym 'a aform = "'a \ 'a pdevs" subsection \Evaluation, Range, Joint Range\ definition aform_val :: "(nat \ real) \ 'a::real_normed_vector aform \ 'a" where "aform_val e X = fst X + pdevs_val e (snd X)" definition Affine :: "'a::real_normed_vector aform \ 'a set" where "Affine X = valuate (\e. aform_val e X)" definition Joints :: "'a::real_normed_vector aform list \ 'a list set" where "Joints XS = valuate (\e. map (aform_val e) XS)" lemma Joints_nthE: assumes "zs \ Joints ZS" obtains e where "\i. i < length zs \ zs ! i = aform_val e (ZS ! i)" "\i. e i \ {-1..1}" using assms by atomize_elim (auto simp: Joints_def Pi_iff valuate_ex) lemma Joints_mapE: assumes "ys \ Joints YS" obtains e where "ys = map (\x. aform_val e x) YS" "\i. e i \ {-1 .. 1}" using assms by (force simp: Joints_def valuate_def) lemma zipped_subset_mapped_Elem: assumes "xs = map (aform_val e) XS" assumes e: "\i. e i \ {-1 .. 1}" assumes [simp]: "length xs = length XS" assumes [simp]: "length ys = length YS" assumes "set (zip ys YS) \ set (zip xs XS)" shows "ys = map (aform_val e) YS" proof - from assms have ys: "\i. i < length xs \ xs ! i = aform_val e (XS ! i)" by auto from assms have set_eq: "{(ys ! i, YS ! i) |i. i < length ys \ i < length YS} \ {(xs ! i, XS ! i) |i. i < length xs \ i < length XS}" using assms(2) by (auto simp: set_zip) hence "\ij YS ! i = XS ! j" by auto then obtain j where j: "\i. i < length YS \ ys ! i = xs ! (j i)" "\i. i < length YS \ YS ! i = XS ! (j i)" "\i. i < length YS \ j i < length XS" by metis show ?thesis using assms by (auto simp: Joints_def j ys intro!: exI[where x=e] nth_equalityI) qed lemma Joints_set_zip_subset: assumes "xs \ Joints XS" assumes "length xs = length XS" assumes "length ys = length YS" assumes "set (zip ys YS) \ set (zip xs XS)" shows "ys \ Joints YS" proof - from Joints_mapE assms obtain e where ys: "xs = map (\x. aform_val e x) XS" and e: "\i. e i \ {-1 .. 1}" by blast show "ys \ Joints YS" using e zipped_subset_mapped_Elem[OF ys e assms(2-4)] by (auto simp: Joints_def valuate_def intro!: exI[where x=e]) qed lemma Joints_set_zip: assumes "ys \ Joints YS" assumes "length xs = length XS" assumes "length YS = length XS" assumes sets_eq: "set (zip xs XS) = set (zip ys YS)" shows "xs \ Joints XS" proof - from assms have "length ys = length YS" by (auto simp: Joints_def valuate_def) from assms(1) this assms(2) show ?thesis by (rule Joints_set_zip_subset) (simp add: assms) qed definition Joints2 :: "'a::real_normed_vector aform list \'b::real_normed_vector aform \ ('a list \ 'b) set" where "Joints2 XS Y = valuate (\e. (map (aform_val e) XS, aform_val e Y))" lemma Joints2E: assumes "zs_y \ Joints2 ZS Y" obtains e where "\i. i < length (fst zs_y) \ (fst zs_y) ! i = aform_val e (ZS ! i)" "snd (zs_y) = aform_val e Y" "\i. e i \ {-1..1}" using assms by atomize_elim (auto simp: Joints2_def Pi_iff valuate_ex) lemma nth_in_AffineI: assumes "xs \ Joints XS" assumes "i < length XS" shows "xs ! i \ Affine (XS ! i)" using assms by (force simp: Affine_def Joints_def valuate_def) lemma Cons_nth_in_Joints1: assumes "xs \ Joints XS" assumes "i < length XS" shows "((xs ! i) # xs) \ Joints ((XS ! i) # XS)" using assms by (force simp: Joints_def valuate_def) lemma Cons_nth_in_Joints2: assumes "xs \ Joints XS" assumes "i < length XS" assumes "j < length XS" shows "((xs ! i) #(xs ! j) # xs) \ Joints ((XS ! i)#(XS ! j) # XS)" using assms by (force simp: Joints_def valuate_def) lemma Joints_swap: "x#y#xs\Joints (X#Y#XS) \ y#x#xs \ Joints (Y#X#XS)" by (force simp: Joints_def valuate_def) lemma Joints_swap_Cons_append: "length xs = length XS \ x#ys@xs\Joints (X#YS@XS) \ ys@x#xs \ Joints (YS@X#XS)" by (auto simp: Joints_def valuate_def) lemma Joints_ConsD: "x#xs\Joints (X#XS) \ xs \ Joints XS" by (force simp: Joints_def valuate_def) lemma Joints_appendD1: "ys@xs\Joints (YS@XS) \ length xs = length XS \ xs \ Joints XS" by (force simp: Joints_def valuate_def) lemma Joints_appendD2: "ys@xs\Joints (YS@XS) \ length ys = length YS \ ys \ Joints YS" by (force simp: Joints_def valuate_def) lemma Joints_imp_length_eq: "xs \ Joints XS \ length xs = length XS" by (auto simp: Joints_def valuate_def) lemma Joints_rotate[simp]: "xs@[x] \ Joints (XS @[X]) \ x#xs \ Joints (X#XS)" by (auto simp: Joints_def valuate_def) subsection \Domain\ definition "pdevs_domain x = {i. pdevs_apply x i \ 0}" lemma finite_pdevs_domain[intro, simp]: "finite (pdevs_domain x)" unfolding pdevs_domain_def by transfer lemma in_pdevs_domain[simp]: "i \ pdevs_domain x \ pdevs_apply x i \ 0" by (auto simp: pdevs_domain_def) subsection \Least Fresh Index\ definition degree::"'a::real_vector pdevs \ nat" where "degree x = (LEAST i. \j\i. pdevs_apply x j = 0)" lemma degree[rule_format, intro, simp]: shows "\j\degree x. pdevs_apply x j = 0" unfolding degree_def proof (rule LeastI_ex) have "\j. j > Max (pdevs_domain x) \ j \ (pdevs_domain x)" by (metis Max_less_iff all_not_in_conv less_irrefl_nat finite_pdevs_domain) then show "\xa. \j\xa. pdevs_apply x j = 0" by (auto intro!: exI[where x="Max (pdevs_domain x) + 1"]) qed lemma degree_le: assumes d: "\j \ d. pdevs_apply x j = 0" shows "degree x \ d" unfolding degree_def by (rule Least_le) (rule d) lemma degree_gt: "pdevs_apply x j \ 0 \ degree x > j" by auto lemma pdevs_val_pdevs_domain: "pdevs_val e X = (\i\pdevs_domain X. e i *\<^sub>R pdevs_apply X i)" by (auto simp: pdevs_val_def intro!: suminf_finite) lemma pdevs_val_sum_le: "degree X \ d \ pdevs_val e X = (\i < d. e i *\<^sub>R pdevs_apply X i)" by (force intro!: degree_gt sum.mono_neutral_cong_left simp: pdevs_val_pdevs_domain) lemmas pdevs_val_sum = pdevs_val_sum_le[OF order_refl] lemma pdevs_val_zero[simp]: "pdevs_val (\_. 0) x = 0" by (auto simp: pdevs_val_sum) lemma degree_eqI: assumes "pdevs_apply x d \ 0" assumes "\j. j > d \ pdevs_apply x j = 0" shows "degree x = Suc d" unfolding eq_iff by (auto intro!: degree_gt degree_le assms simp: Suc_le_eq) lemma finite_degree_nonzero[intro, simp]: "finite {i. pdevs_apply x i \ 0}" by transfer (auto simp: vimage_def Collect_neg_eq) lemma degree_eq_Suc_max: "degree x = (if (\i. pdevs_apply x i = 0) then 0 else Suc (Max {i. pdevs_apply x i \ 0}))" proof - { assume "\i. pdevs_apply x i = 0" hence ?thesis by auto (metis degree_le le_0_eq) } moreover { fix i assume "pdevs_apply x i \ 0" hence ?thesis using Max_in[OF finite_degree_nonzero, of x] by (auto intro!: degree_eqI) (metis Max.coboundedI[OF finite_degree_nonzero] in_pdevs_domain le_eq_less_or_eq less_asym pdevs_domain_def) } ultimately show ?thesis by blast qed lemma pdevs_val_degree_cong: assumes "b = d" assumes "\i. i < degree b \ a i = c i" shows "pdevs_val a b = pdevs_val c d" using assms by (auto simp: pdevs_val_sum) abbreviation degree_aform::"'a::real_vector aform \ nat" where "degree_aform X \ degree (snd X)" lemma degree_cong: "(\i. (pdevs_apply x i = 0) = (pdevs_apply y i = 0)) \ degree x = degree y" unfolding degree_def by auto lemma Least_True_nat[intro, simp]: "(LEAST i::nat. True) = 0" by (metis (lifting) One_nat_def less_one not_less_Least not_less_eq) lemma sorted_list_of_pdevs_domain_eq: "sorted_list_of_set (pdevs_domain X) = filter ((\) 0 o pdevs_apply X) [0..x. x", simplified]) subsection \Total Deviation\ definition tdev::"'a::ordered_euclidean_space pdevs \ 'a" where "tdev x = (\ipdevs_apply x i\)" lemma abs_pdevs_val_le_tdev: "e \ UNIV \ {-1 .. 1} \ \pdevs_val e x\ \ tdev x" by (force simp: pdevs_val_sum tdev_def abs_scaleR Pi_iff intro!: order_trans[OF sum_abs] sum_mono scaleR_left_le_one_le intro: abs_leI) subsection \Binary Pointwise Operations\ definition binop_pdevs_raw::"('a::zero \ 'b::zero \ 'c::zero) \ (nat \ 'a) \ (nat \ 'b) \ nat \ 'c" where "binop_pdevs_raw f x y i = (if x i = 0 \ y i = 0 then 0 else f (x i) (y i))" lemma nonzeros_binop_pdevs_subset: "{i. binop_pdevs_raw f x y i \ 0} \ {i. x i \ 0} \ {i. y i \ 0}" by (auto simp: binop_pdevs_raw_def) lift_definition binop_pdevs:: "('a \ 'b \ 'c) \ 'a::zero pdevs \ 'b::zero pdevs \ 'c::zero pdevs" is binop_pdevs_raw using nonzeros_binop_pdevs_subset by (rule finite_subset) auto lemma pdevs_apply_binop_pdevs[simp]: "pdevs_apply (binop_pdevs f x y) i = (if pdevs_apply x i = 0 \ pdevs_apply y i = 0 then 0 else f (pdevs_apply x i) (pdevs_apply y i))" by transfer (auto simp: binop_pdevs_raw_def) subsection \Addition\ definition add_pdevs::"'a::real_vector pdevs \ 'a pdevs \ 'a pdevs" where "add_pdevs = binop_pdevs (+)" lemma pdevs_apply_add_pdevs[simp]: "pdevs_apply (add_pdevs X Y) n = pdevs_apply X n + pdevs_apply Y n" by (auto simp: add_pdevs_def) lemma pdevs_val_add_pdevs[simp]: fixes x y::"'a::euclidean_space" shows "pdevs_val e (add_pdevs X Y) = pdevs_val e X + pdevs_val e Y" proof - let ?sum = "\m X. \i < m. e i *\<^sub>R pdevs_apply X i" let ?m = "max (degree X) (degree Y)" have "pdevs_val e X + pdevs_val e Y = ?sum (degree X) X + ?sum (degree Y) Y" by (simp add: pdevs_val_sum) also have "?sum (degree X) X = ?sum ?m X" by (rule sum.mono_neutral_cong_left) auto also have "?sum (degree Y) Y = ?sum ?m Y" by (rule sum.mono_neutral_cong_left) auto also have "?sum ?m X + ?sum ?m Y = (\i < ?m. e i *\<^sub>R (pdevs_apply X i + pdevs_apply Y i))" by (simp add: scaleR_right_distrib sum.distrib) also have "\ = (\i. e i *\<^sub>R (pdevs_apply X i + pdevs_apply Y i))" by (rule suminf_finite[symmetric]) auto also have "\ = pdevs_val e (add_pdevs X Y)" by (simp add: pdevs_val_def) finally show "pdevs_val e (add_pdevs X Y) = pdevs_val e X + pdevs_val e Y" by simp qed subsection \Total Deviation\ lemma tdev_eq_zero_iff: fixes X::"real pdevs" shows "tdev X = 0 \ (\e. pdevs_val e X = 0)" by (force simp add: pdevs_val_sum tdev_def sum_nonneg_eq_0_iff dest!: spec[where x="\i. if pdevs_apply X i \ 0 then 1 else -1"] split: if_split_asm) lemma tdev_nonneg[intro, simp]: "tdev X \ 0" by (auto simp: tdev_def) lemma tdev_nonpos_iff[simp]: "tdev X \ 0 \ tdev X = 0" by (auto simp: order.antisym) subsection \Unary Operations\ definition unop_pdevs_raw:: "('a::zero \ 'b::zero) \ (nat \ 'a) \ nat \ 'b" where "unop_pdevs_raw f x i = (if x i = 0 then 0 else f (x i))" lemma nonzeros_unop_pdevs_subset: "{i. unop_pdevs_raw f x i \ 0} \ {i. x i \ 0}" by (auto simp: unop_pdevs_raw_def) lift_definition unop_pdevs:: "('a \ 'b) \ 'a::zero pdevs \ 'b::zero pdevs" is unop_pdevs_raw using nonzeros_unop_pdevs_subset by (rule finite_subset) auto lemma pdevs_apply_unop_pdevs[simp]: "pdevs_apply (unop_pdevs f x) i = (if pdevs_apply x i = 0 then 0 else f (pdevs_apply x i))" by transfer (auto simp: unop_pdevs_raw_def) lemma pdevs_domain_unop_linear: assumes "linear f" shows "pdevs_domain (unop_pdevs f x) \ pdevs_domain x" proof - interpret f: linear f by fact show ?thesis by (auto simp: f.zero) qed lemma pdevs_val_unop_linear: assumes "linear f" shows "pdevs_val e (unop_pdevs f x) = f (pdevs_val e x)" proof - interpret f: linear f by fact have *: "\i. (if pdevs_apply x i = 0 then 0 else f (pdevs_apply x i)) = f (pdevs_apply x i)" by (auto simp: f.zero) have "pdevs_val e (unop_pdevs f x) = (\i\pdevs_domain (unop_pdevs f x). e i *\<^sub>R f (pdevs_apply x i))" by (auto simp add: pdevs_val_pdevs_domain *) also have "\ = (\xa\pdevs_domain x. e xa *\<^sub>R f (pdevs_apply x xa))" by (auto intro!: sum.mono_neutral_cong_left) also have "\ = f (pdevs_val e x)" by (auto simp add: pdevs_val_pdevs_domain f.sum f.scaleR) finally show ?thesis . qed subsection \Pointwise Scaling of Partial Deviations\ definition scaleR_pdevs::"real \ 'a::real_vector pdevs \ 'a pdevs" where "scaleR_pdevs r x = unop_pdevs ((*\<^sub>R) r) x" lemma pdevs_apply_scaleR_pdevs[simp]: "pdevs_apply (scaleR_pdevs x Y) n = x *\<^sub>R pdevs_apply Y n" by (auto simp: scaleR_pdevs_def) lemma degree_scaleR_pdevs[simp]: "degree (scaleR_pdevs r x) = (if r = 0 then 0 else degree x)" unfolding degree_def by auto lemma pdevs_val_scaleR_pdevs[simp]: fixes x::real and Y::"'a::real_normed_vector pdevs" shows "pdevs_val e (scaleR_pdevs x Y) = x *\<^sub>R pdevs_val e Y" by (auto simp: pdevs_val_sum scaleR_sum_right ac_simps) subsection \Partial Deviations Scale Pointwise\ definition pdevs_scaleR::"real pdevs \ 'a::real_vector \ 'a pdevs" where "pdevs_scaleR r x = unop_pdevs (\r. r *\<^sub>R x) r" lemma pdevs_apply_pdevs_scaleR[simp]: "pdevs_apply (pdevs_scaleR X y) n = pdevs_apply X n *\<^sub>R y" by (auto simp: pdevs_scaleR_def) lemma degree_pdevs_scaleR[simp]: "degree (pdevs_scaleR r x) = (if x = 0 then 0 else degree r)" unfolding degree_def by auto lemma pdevs_val_pdevs_scaleR[simp]: fixes X::"real pdevs" and y::"'a::real_normed_vector" shows "pdevs_val e (pdevs_scaleR X y) = pdevs_val e X *\<^sub>R y" by (auto simp: pdevs_val_sum scaleR_sum_left) subsection \Pointwise Unary Minus\ definition uminus_pdevs::"'a::real_vector pdevs \ 'a pdevs" where "uminus_pdevs = unop_pdevs uminus" lemma pdevs_apply_uminus_pdevs[simp]: "pdevs_apply (uminus_pdevs x) = - pdevs_apply x" by (auto simp: uminus_pdevs_def) lemma degree_uminus_pdevs[simp]: "degree (uminus_pdevs x) = degree x" by (rule degree_cong) simp lemma pdevs_val_uminus_pdevs[simp]: "pdevs_val e (uminus_pdevs x) = - pdevs_val e x" unfolding pdevs_val_sum by (auto simp: sum_negf) definition "uminus_aform X = (- fst X, uminus_pdevs (snd X))" lemma fst_uminus_aform[simp]: "fst (uminus_aform Y) = - fst Y" by (simp add: uminus_aform_def) lemma aform_val_uminus_aform[simp]: "aform_val e (uminus_aform X) = - aform_val e X" by (auto simp: uminus_aform_def aform_val_def) subsection \Constant\ lift_definition zero_pdevs::"'a::zero pdevs" is "\_. 0" by simp lemma pdevs_apply_zero_pdevs[simp]: "pdevs_apply zero_pdevs i = 0" by transfer simp lemma pdevs_val_zero_pdevs[simp]: "pdevs_val e zero_pdevs = 0" by (auto simp: pdevs_val_def) definition "num_aform f = (f, zero_pdevs)" subsection \Inner Product\ definition pdevs_inner::"'a::euclidean_space pdevs \ 'a \ real pdevs" where "pdevs_inner x b = unop_pdevs (\x. x \ b) x" lemma pdevs_apply_pdevs_inner[simp]: "pdevs_apply (pdevs_inner p a) i = pdevs_apply p i \ a" by (simp add: pdevs_inner_def) lemma pdevs_val_pdevs_inner[simp]: "pdevs_val e (pdevs_inner p a) = pdevs_val e p \ a" by (auto simp add: inner_sum_left pdevs_val_pdevs_domain intro!: sum.mono_neutral_cong_left) definition inner_aform::"'a::euclidean_space aform \ 'a \ real aform" where "inner_aform X b = (fst X \ b, pdevs_inner (snd X) b)" subsection \Inner Product Pair\ definition inner2::"'a::euclidean_space \ 'a \ 'a \ real*real" where "inner2 x n l = (x \ n, x \ l)" definition pdevs_inner2::"'a::euclidean_space pdevs \ 'a \ 'a \ (real*real) pdevs" where "pdevs_inner2 X n l = unop_pdevs (\x. inner2 x n l) X" lemma pdevs_apply_pdevs_inner2[simp]: "pdevs_apply (pdevs_inner2 p a b) i = (pdevs_apply p i \ a, pdevs_apply p i \ b)" by (simp add: pdevs_inner2_def inner2_def zero_prod_def) definition inner2_aform::"'a::euclidean_space aform \ 'a \ 'a \ (real*real) aform" where "inner2_aform X a b = (inner2 (fst X) a b, pdevs_inner2 (snd X) a b)" lemma linear_inner2[intro, simp]: "linear (\x. inner2 x n i)" by unfold_locales (auto simp: inner2_def algebra_simps) lemma aform_val_inner2_aform[simp]: "aform_val e (inner2_aform Z n i) = inner2 (aform_val e Z) n i" proof - have "aform_val e (inner2_aform Z n i) = inner2 (fst Z) n i + inner2 (pdevs_val e (snd Z)) n i" by (auto simp: aform_val_def inner2_aform_def pdevs_inner2_def pdevs_val_unop_linear) also have "\ = inner2 (aform_val e Z) n i" by (simp add: inner2_def algebra_simps aform_val_def) finally show ?thesis . qed subsection \Update\ lemma pdevs_val_upd[simp]: "pdevs_val (e(n := e')) X = pdevs_val e X - e n * pdevs_apply X n + e' * pdevs_apply X n" unfolding pdevs_val_def by (subst suminf_finite[OF finite.insertI[OF finite_degree_nonzero], of n X], auto simp: pdevs_val_def sum.insert_remove)+ lemma nonzeros_fun_upd: "{i. (f(n := a)) i \ 0} \ {i. f i \ 0} \ {n}" by (auto split: if_split_asm) lift_definition pdev_upd::"'a::real_vector pdevs \ nat \ 'a \ 'a pdevs" is "\x n a. x(n:=a)" by (rule finite_subset[OF nonzeros_fun_upd]) simp lemma pdevs_apply_pdev_upd[simp]: "pdevs_apply (pdev_upd X n x) = (pdevs_apply X)(n:=x)" by transfer simp lemma pdevs_val_pdev_upd[simp]: "pdevs_val e (pdev_upd X n x) = pdevs_val e X + e n *\<^sub>R x - e n *\<^sub>R pdevs_apply X n" unfolding pdevs_val_def by (subst suminf_finite[OF finite.insertI[OF finite_degree_nonzero], of n X], auto simp: pdevs_val_def sum.insert_remove)+ lemma degree_pdev_upd: assumes "x = 0 \ pdevs_apply X n = 0" shows "degree (pdev_upd X n x) = degree X" using assms by (auto intro!: degree_cong split: if_split_asm) lemma degree_pdev_upd_le: assumes "degree X \ n" shows "degree (pdev_upd X n x) \ Suc n" using assms by (auto intro!: degree_le) subsection \Inf/Sup\ definition "Inf_aform X = fst X - tdev (snd X)" definition "Sup_aform X = fst X + tdev (snd X)" lemma Inf_aform: assumes "e \ UNIV \ {-1 .. 1}" shows "Inf_aform X \ aform_val e X" using order_trans[OF abs_ge_minus_self abs_pdevs_val_le_tdev[OF assms]] by (auto simp: Inf_aform_def aform_val_def minus_le_iff) lemma Sup_aform: assumes "e \ UNIV \ {-1 .. 1}" shows "aform_val e X \ Sup_aform X" using order_trans[OF abs_ge_self abs_pdevs_val_le_tdev[OF assms]] by (auto simp: Sup_aform_def aform_val_def) subsection \Minkowski Sum\ definition msum_pdevs_raw::"nat\(nat \ 'a::real_vector)\(nat \ 'a)\(nat\'a)" where "msum_pdevs_raw n x y i = (if i < n then x i else y (i - n))" lemma nonzeros_msum_pdevs_raw: "{i. msum_pdevs_raw n f g i \ 0} = ({0.. {i. f i \ 0}) \ (+) n ` ({i. g i \ 0})" by (force simp: msum_pdevs_raw_def not_less split: if_split_asm) lift_definition msum_pdevs::"nat\'a::real_vector pdevs\'a pdevs\'a pdevs" is msum_pdevs_raw unfolding nonzeros_msum_pdevs_raw by simp lemma pdevs_apply_msum_pdevs: "pdevs_apply (msum_pdevs n f g) i = (if i < n then pdevs_apply f i else pdevs_apply g (i - n))" by transfer (auto simp: msum_pdevs_raw_def) lemma degree_least_nonzero: assumes "degree f \ 0" shows "pdevs_apply f (degree f - 1) \ 0" proof assume H: "pdevs_apply f (degree f - 1) = 0" { fix j assume "j\degree f - 1" with H have "pdevs_apply f j = 0" by (cases "degree f - 1 = j") auto } from degree_le[rule_format, OF this] have "degree f \ degree f - 1" by blast with assms show False by simp qed lemma degree_leI: assumes "(\i. pdevs_apply y i = 0 \ pdevs_apply x i = 0)" shows "degree x \ degree y" proof cases assume "degree x \ 0" from degree_least_nonzero[OF this] have "pdevs_apply y (degree x - 1) \ 0" by (auto simp: assms split: if_split_asm) from degree_gt[OF this] show ?thesis by simp qed simp lemma degree_msum_pdevs_ge1: shows "degree f \ n \ degree f \ degree (msum_pdevs n f g)" by (rule degree_leI) (auto simp: pdevs_apply_msum_pdevs split: if_split_asm) lemma degree_msum_pdevs_ge2: assumes "degree f \ n" shows "degree g \ degree (msum_pdevs n f g) - n" proof cases assume "degree g \ 0" hence "pdevs_apply g (degree g - 1) \ 0" by (rule degree_least_nonzero) hence "pdevs_apply (msum_pdevs n f g) (n + degree g - 1) \ 0" using assms by (auto simp: pdevs_apply_msum_pdevs) from degree_gt[OF this] show ?thesis by simp qed simp lemma degree_msum_pdevs_le: shows "degree (msum_pdevs n f g) \ n + degree g" by (auto intro!: degree_le simp: pdevs_apply_msum_pdevs) lemma sum_msum_pdevs_cases: assumes "degree f \ n" assumes [simp]: "\i. e i 0 = 0" shows "(\i i i i\{.. {i. i < n}. e i (pdevs_apply f i)) + (\i\{.. - {i. i < n}. e i (pdevs_apply g (i - n)))" (is "_ = ?sum_f + ?sum_g") by (simp add: sum.If_cases if_distrib) also have "?sum_f = (\i = 0..i\{0 + n.. = (\i = 0.. = (\i = 0.. n \ tdev (msum_pdevs n f g) = tdev f + tdev g" by (auto simp: tdev_def pdevs_apply_msum_pdevs intro!: sum_msum_pdevs_cases) lemma pdevs_val_msum_pdevs: "degree f \ n \ pdevs_val e (msum_pdevs n f g) = pdevs_val e f + pdevs_val (\i. e (i + n)) g" by (auto simp: pdevs_val_sum pdevs_apply_msum_pdevs intro!: sum_msum_pdevs_cases) definition msum_aform::"nat \ 'a::real_vector aform \ 'a aform \ 'a aform" where "msum_aform n f g = (fst f + fst g, msum_pdevs n (snd f) (snd g))" lemma fst_msum_aform[simp]: "fst (msum_aform n f g) = fst f + fst g" by (simp add: msum_aform_def) lemma snd_msum_aform[simp]: "snd (msum_aform n f g) = msum_pdevs n (snd f) (snd g)" by (simp add: msum_aform_def) lemma finite_nonzero_summable: "finite {i. f i \ 0} \ summable f" by (auto intro!: sums_summable sums_finite) lemma aform_val_msum_aform: assumes "degree_aform f \ n" shows "aform_val e (msum_aform n f g) = aform_val e f + aform_val (\i. e (i + n)) g" using assms by (auto simp: pdevs_val_msum_pdevs aform_val_def) lemma Inf_aform_msum_aform: "degree_aform X \ n \ Inf_aform (msum_aform n X Y) = Inf_aform X + Inf_aform Y" by (simp add: Inf_aform_def tdev_msum_pdevs) lemma Sup_aform_msum_aform: "degree_aform X \ n \ Sup_aform (msum_aform n X Y) = Sup_aform X + Sup_aform Y" by (simp add: Sup_aform_def tdev_msum_pdevs) definition "independent_from d Y = msum_aform d (0, zero_pdevs) Y" definition "independent_aform X Y = independent_from (degree_aform X) Y" lemma degree_zero_pdevs[simp]: "degree zero_pdevs = 0" by (metis degree_least_nonzero pdevs_apply_zero_pdevs) lemma independent_aform_Joints: assumes "x \ Affine X" assumes "y \ Affine Y" shows "[x, y] \ Joints [X, independent_aform X Y]" using assms unfolding Affine_def valuate_def Joints_def apply safe subgoal premises prems for e ea using prems by (intro image_eqI[where x="\i. if i < degree_aform X then e i else ea (i - degree_aform X)"]) (auto simp: aform_val_def pdevs_val_msum_pdevs Pi_iff independent_aform_def independent_from_def intro!: pdevs_val_degree_cong) done lemma msum_aform_Joints: assumes "d \ degree_aform X" assumes "\X. X \ set XS \ d \ degree_aform X" assumes "(x#xs) \ Joints (X#XS)" assumes "y \ Affine Y" shows "((x + y)#x#xs) \ Joints (msum_aform d X Y#X#XS)" using assms unfolding Joints_def valuate_def Affine_def proof (safe, goal_cases) case (1 e ea a b zs) then show ?case by (intro image_eqI[where x = "\i. if i < d then e i else ea (i - d)"]) (force simp: aform_val_def pdevs_val_msum_pdevs intro!: intro!: pdevs_val_degree_cong)+ qed lemma Joints_msum_aform: assumes "d \ degree_aform X" assumes "\Y. Y \ set YS \ d \ degree_aform Y" shows "Joints (msum_aform d X Y#YS) = {((x + y)#ys) |x y ys. y \ Affine Y \ x#ys \ Joints (X#YS)}" unfolding Affine_def valuate_def Joints_def proof (safe, goal_cases) case (1 x e) thus ?case using assms by (intro exI[where x = "aform_val e X"] exI[where x = "aform_val ((\i. e (i + d))) Y"]) (auto simp add: aform_val_def pdevs_val_msum_pdevs) next case (2 x xa y ys e ea) thus ?case using assms by (intro image_eqI[where x="\i. if i < d then ea i else e (i - d)"]) (force simp: aform_val_def pdevs_val_msum_pdevs Pi_iff intro!: pdevs_val_degree_cong)+ qed lemma Joints_singleton_image: "Joints [x] = (\x. [x]) ` Affine x" by (auto simp: Joints_def Affine_def valuate_def) lemma Collect_extract_image: "{g (f x y) |x y. P x y} = g ` {f x y |x y. P x y}" by auto lemma inj_Cons: "inj (\x. x#xs)" by (auto intro!: injI) lemma Joints_Nil[simp]: "Joints [] = {[]}" by (force simp: Joints_def valuate_def) lemma msum_pdevs_zero_ident[simp]: "msum_pdevs 0 zero_pdevs x = x" by transfer (auto simp: msum_pdevs_raw_def) lemma msum_aform_zero_ident[simp]: "msum_aform 0 (0, zero_pdevs) x = x" by (simp add: msum_aform_def) lemma mem_Joints_singleton: "(x \ Joints [X]) = (\y. x = [y] \ y \ Affine X)" by (auto simp: Affine_def valuate_def Joints_def) lemma singleton_mem_Joints[simp]: "[x] \ Joints [X] \ x \ Affine X" by (auto simp: mem_Joints_singleton) lemma msum_aform_Joints_without_first: assumes "d \ degree_aform X" assumes "\X. X \ set XS \ d \ degree_aform X" assumes "(x#xs) \ Joints (X#XS)" assumes "y \ Affine Y" assumes "z = x + y" shows "z#xs \ Joints (msum_aform d X Y#XS)" unfolding \z = x + y\ using msum_aform_Joints[OF assms(1-4)] by (force simp: Joints_def valuate_def) lemma Affine_msum_aform: assumes "d \ degree_aform X" shows "Affine (msum_aform d X Y) = {x + y |x y. x \ Affine X \ y \ Affine Y}" using Joints_msum_aform[OF assms, of Nil Y, simplified, unfolded mem_Joints_singleton] by (auto simp add: Joints_singleton_image Collect_extract_image[where g="\x. [x]"] inj_image_eq_iff[OF inj_Cons] ) lemma Affine_zero_pdevs[simp]: "Affine (0, zero_pdevs) = {0}" by (force simp: Affine_def valuate_def aform_val_def) lemma Affine_independent_aform: "Affine (independent_aform X Y) = Affine Y" by (auto simp: independent_aform_def independent_from_def Affine_msum_aform) lemma abs_diff_eq1: fixes l u::"'a::ordered_euclidean_space" shows "l \ u \ \u - l\ = u - l" by (metis abs_of_nonneg diff_add_cancel le_add_same_cancel2) lemma compact_sum: fixes f :: "'a \ 'b::topological_space \ 'c::real_normed_vector" assumes "finite I" assumes "\i. i \ I \ compact (S i)" assumes "\i. i \ I \ continuous_on (S i) (f i)" assumes "I \ J" shows "compact {\i\I. f i (x i) | x. x \ Pi J S}" using assms proof (induct I) case empty thus ?case proof (cases "\x. x \ Pi J S") case False hence *: "{\i\{}. f i (x i) |x. x \ Pi J S} = {}" by (auto simp: Pi_iff) show ?thesis unfolding * by simp qed auto next case (insert a I) hence "{\i\insert a I. f i (xa i) |xa. xa \ Pi J S} = {x + y |x y. x \ f a ` S a \ y \ {\i\I. f i (x i) |x. x \ Pi J S}}" proof safe fix s x assume "s \ S a" "x \ Pi J S" thus "\xa. f a s + (\i\I. f i (x i)) = (\i\insert a I. f i (xa i)) \ xa \ Pi J S" using insert by (auto intro!: exI[where x="x(a:=s)"] sum.cong) qed force also have "compact \" using insert by (intro compact_sums) (auto intro!: compact_continuous_image) finally show ?case . qed lemma compact_Affine: fixes X::"'a::ordered_euclidean_space aform" shows "compact (Affine X)" proof - have "Affine X = {x + y|x y. x \ {fst X} \ y \ {(\i \ {0..R pdevs_apply (snd X) i) | e. e \ UNIV \ {-1 .. 1}}}" by (auto simp: Affine_def valuate_def aform_val_def pdevs_val_sum atLeast0LessThan) also have "compact \" by (rule compact_sums) (auto intro!: compact_sum continuous_intros) finally show ?thesis . qed lemma Joints2_JointsI: "(xs, x) \ Joints2 XS X \ x#xs \ Joints (X#XS)" by (auto simp: Joints_def Joints2_def valuate_def) subsection \Splitting\ definition "split_aform X i = (let xi = pdevs_apply (snd X) i /\<^sub>R 2 in ((fst X - xi, pdev_upd (snd X) i xi), (fst X + xi, pdev_upd (snd X) i xi)))" lemma split_aformE: assumes "e \ UNIV \ {-1 .. 1}" assumes "x = aform_val e X" obtains err where "x = aform_val (e(i:=err)) (fst (split_aform X i))" "err \ {-1 .. 1}" | err where "x = aform_val (e(i:=err)) (snd (split_aform X i))" "err \ {-1 .. 1}" proof (atomize_elim) let ?thesis = "(\err. x = aform_val (e(i := err)) (fst (split_aform X i)) \ err \ {- 1..1}) \ (\err. x = aform_val (e(i := err)) (snd (split_aform X i)) \ err \ {- 1..1})" { assume "pdevs_apply (snd X) i = 0" hence "X = fst (split_aform X i)" by (auto simp: split_aform_def intro!: prod_eqI pdevs_eqI) with assms have ?thesis by (auto intro!: exI[where x="e i"]) } moreover { assume "pdevs_apply (snd X) i \ 0" hence [simp]: "degree_aform X > i" by (rule degree_gt) note assms(2) also have "aform_val e X = fst X + (\iR pdevs_apply (snd X) i)" by (simp add: aform_val_def pdevs_val_sum) also have rewr: "{.. {i}" by auto have "(\iR pdevs_apply (snd X) i) = (\i \ {0..R pdevs_apply (snd X) i) + e i *\<^sub>R pdevs_apply (snd X) i" by (subst rewr, subst sum.union_disjoint) auto finally have "x = fst X + \" . hence "x = aform_val (e(i:=2 * e i - 1)) (snd (split_aform X i))" "x = aform_val (e(i:=2 * e i + 1)) (fst (split_aform X i))" by (auto simp: aform_val_def split_aform_def Let_def pdevs_val_sum atLeast0LessThan Diff_eq degree_pdev_upd if_distrib sum.If_cases field_simps scaleR_left_distrib[symmetric]) moreover have "2 * e i - 1 \ {-1 .. 1} \ 2 * e i + 1 \ {-1 .. 1}" using assms by (auto simp: not_le Pi_iff dest!: spec[where x=i]) ultimately have ?thesis by blast } ultimately show ?thesis by blast qed lemma pdevs_val_add: "pdevs_val (\i. e i + f i) xs = pdevs_val e xs + pdevs_val f xs" by (auto simp: pdevs_val_pdevs_domain algebra_simps sum.distrib) lemma pdevs_val_minus: "pdevs_val (\i. e i - f i) xs = pdevs_val e xs - pdevs_val f xs" by (auto simp: pdevs_val_pdevs_domain algebra_simps sum_subtractf) lemma pdevs_val_cmul: "pdevs_val (\i. u * e i) xs = u *\<^sub>R pdevs_val e xs" by (auto simp: pdevs_val_pdevs_domain scaleR_sum_right) lemma atLeastAtMost_absI: "- a \ a \ \x::real\ \ \a\ \ x \ atLeastAtMost (- a) a" by auto lemma divide_atLeastAtMost_1_absI: "\x::real\ \ \a\ \ x/a \ {-1 .. 1}" by (intro atLeastAtMost_absI) (auto simp: divide_le_eq_1) lemma convex_scaleR_aux: "u + v = 1 \ u *\<^sub>R x + v *\<^sub>R x = (x::'a::real_vector)" by (metis scaleR_add_left scaleR_one) lemma convex_mult_aux: "u + v = 1 \ u * x + v * x = (x::real)" using convex_scaleR_aux[of u v x] by simp lemma convex_Affine: "convex (Affine X)" proof (rule convexI) fix x y::'a and u v::real assume "x \ Affine X" "y \ Affine X" and convex: "0 \ u" "0 \ v" "u + v = 1" then obtain e f where x: "x = aform_val e X" "e \ UNIV \ {-1 .. 1}" and y: "y = aform_val f X" "f \ UNIV \ {-1 .. 1}" by (auto simp: Affine_def valuate_def) let ?conv = "\i. u * e i + v * f i" { fix i have "\?conv i\ \ u * \e i\ + v * \f i\" using convex by (intro order_trans[OF abs_triangle_ineq]) (simp add: abs_mult) also have "\ \ 1" using convex x y by (intro convex_bound_le) (auto simp: Pi_iff abs_real_def) finally have "?conv i \ 1" "-1 \ ?conv i" by (auto simp: abs_real_def split: if_split_asm) } thus "u *\<^sub>R x + v *\<^sub>R y \ Affine X" using convex x y by (auto simp: Affine_def valuate_def aform_val_def pdevs_val_add pdevs_val_cmul algebra_simps convex_scaleR_aux intro!: image_eqI[where x="?conv"]) qed lemma segment_in_aform_val: assumes "e \ UNIV \ {-1 .. 1}" assumes "f \ UNIV \ {-1 .. 1}" shows "closed_segment (aform_val e X) (aform_val f X) \ Affine X" proof - have "aform_val e X \ Affine X" "aform_val f X \ Affine X" using assms by (auto simp: Affine_def valuate_def) with convex_Affine[of X, simplified convex_contains_segment] show ?thesis by simp qed subsection \From List of Generators\ lift_definition pdevs_of_list::"'a::zero list \ 'a pdevs" is "\xs i. if i < length xs then xs ! i else 0" by auto lemma pdevs_apply_pdevs_of_list: "pdevs_apply (pdevs_of_list xs) i = (if i < length xs then xs ! i else 0)" by transfer simp lemma pdevs_apply_pdevs_of_list_Nil[simp]: "pdevs_apply (pdevs_of_list []) i = 0" by transfer auto lemma pdevs_apply_pdevs_of_list_Cons: "pdevs_apply (pdevs_of_list (x # xs)) i = (if i = 0 then x else pdevs_apply (pdevs_of_list xs) (i - 1))" by transfer auto lemma pdevs_domain_pdevs_of_list_Cons[simp]: "pdevs_domain (pdevs_of_list (x # xs)) = (if x = 0 then {} else {0}) \ (+) 1 ` pdevs_domain (pdevs_of_list xs)" by (force simp: pdevs_apply_pdevs_of_list_Cons split: if_split_asm) lemma pdevs_val_pdevs_of_list_eq[simp]: "pdevs_val e (pdevs_of_list (x # xs)) = e 0 *\<^sub>R x + pdevs_val (e o (+) 1) (pdevs_of_list xs)" proof - have "pdevs_val e (pdevs_of_list (x # xs)) = (\i\pdevs_domain (pdevs_of_list (x # xs)) \ {0}. e i *\<^sub>R x) + (\i\pdevs_domain (pdevs_of_list (x # xs)) \ - {0}. e i *\<^sub>R pdevs_apply (pdevs_of_list xs) (i - Suc 0))" (is "_ = ?l + ?r") by (simp add: pdevs_val_pdevs_domain if_distrib sum.If_cases pdevs_apply_pdevs_of_list_Cons) also have "?r = (\i\pdevs_domain (pdevs_of_list xs). e (Suc i) *\<^sub>R pdevs_apply (pdevs_of_list xs) i)" by (rule sum.reindex_cong[of "\i. i + 1"]) auto also have "\ = pdevs_val (e o (+) 1) (pdevs_of_list xs)" by (simp add: pdevs_val_pdevs_domain ) also have "?l = (\i\{0}. e i *\<^sub>R x)" by (rule sum.mono_neutral_cong_left) auto also have "\ = e 0 *\<^sub>R x" by simp finally show ?thesis . qed lemma less_degree_pdevs_of_list_imp_less_length: assumes "i < degree (pdevs_of_list xs)" shows "i < length xs" proof - from assms have "pdevs_apply (pdevs_of_list xs) (degree (pdevs_of_list xs) - 1) \ 0" by (metis degree_least_nonzero less_nat_zero_code) hence "degree (pdevs_of_list xs) - 1 < length xs" by (simp add: pdevs_apply_pdevs_of_list split: if_split_asm) with assms show ?thesis by simp qed lemma tdev_pdevs_of_list[simp]: "tdev (pdevs_of_list xs) = sum_list (map abs xs)" by (auto simp: tdev_def pdevs_apply_pdevs_of_list sum_list_sum_nth less_degree_pdevs_of_list_imp_less_length intro!: sum.mono_neutral_cong_left degree_gt) lemma pdevs_of_list_Nil[simp]: "pdevs_of_list [] = zero_pdevs" by (auto intro!: pdevs_eqI) lemma pdevs_val_inj_sumI: fixes K::"'a set" and g::"'a \ nat" assumes "finite K" assumes "inj_on g K" assumes "pdevs_domain x \ g ` K" assumes "\i. i \ K \ g i \ pdevs_domain x \ f i = 0" assumes "\i. i \ K \ g i \ pdevs_domain x \ f i = e (g i) *\<^sub>R pdevs_apply x (g i)" shows "pdevs_val e x = (\i\K. f i)" proof - have [simp]: "inj_on (the_inv_into K g) (pdevs_domain x)" using assms by (auto simp: intro!: subset_inj_on[OF inj_on_the_inv_into]) { fix y assume y: "y \ pdevs_domain x" have g_inv: "g (the_inv_into K g y) = y" by (meson assms(2) assms(3) y f_the_inv_into_f subset_eq) have inv_in: "the_inv_into K g y \ K" by (meson assms(2) assms(3) y subset_iff in_pdevs_domain the_inv_into_into) have inv3: "the_inv_into (pdevs_domain x) (the_inv_into K g) (the_inv_into K g y) = g (the_inv_into K g y)" using assms y by (subst the_inv_into_f_f) (auto simp: f_the_inv_into_f[OF assms(2)]) note g_inv inv_in inv3 } note this[simp] have "pdevs_val e x = (\i\pdevs_domain x. e i *\<^sub>R pdevs_apply x i)" by (simp add: pdevs_val_pdevs_domain) also have "\ = (\i \ the_inv_into K g ` pdevs_domain x. e (g i) *\<^sub>R pdevs_apply x (g i))" by (rule sum.reindex_cong[OF inj_on_the_inv_into]) auto also have "\ = (\i\K. f i)" using assms by (intro sum.mono_neutral_cong_left) (auto simp: the_inv_into_image_eq) finally show ?thesis . qed lemma pdevs_domain_pdevs_of_list_le: "pdevs_domain (pdevs_of_list xs) \ {0..(i,x)\zip [0..R x)" by (auto simp: sum_list_distinct_conv_sum_set in_set_zip image_fst_zip pdevs_apply_pdevs_of_list distinct_zipI1 intro!: pdevs_val_inj_sumI[of _ fst] split: if_split_asm) lemma scaleR_sum_list: fixes xs::"'a::real_vector list" shows "a *\<^sub>R sum_list xs = sum_list (map (scaleR a) xs)" by (induct xs) (auto simp: algebra_simps) lemma pdevs_val_const_pdevs_of_list: "pdevs_val (\_. c) (pdevs_of_list xs) = c *\<^sub>R sum_list xs" unfolding pdevs_val_zip split_beta' scaleR_sum_list by (rule arg_cong) (auto intro!: nth_equalityI) lemma pdevs_val_partition: assumes "e \ UNIV \ I" obtains f g where "pdevs_val e (pdevs_of_list xs) = pdevs_val f (pdevs_of_list (filter p xs)) + pdevs_val g (pdevs_of_list (filter (Not o p) xs))" "f \ UNIV \ I" "g \ UNIV \ I" proof - obtain i where i: "i \ I" by (metis assms funcset_mem iso_tuple_UNIV_I) let ?zip = "zip [0.. snd) ?zip" let ?f = "(\n. if n < degree (pdevs_of_list (filter p xs)) then e (map fst (fst part) ! n) else i)" let ?g = "(\n. if n < degree (pdevs_of_list (filter (Not \ p) xs)) then e (map fst (snd part) ! n) else i)" show ?thesis proof have "pdevs_val e (pdevs_of_list xs) = (\(i,x)\?zip. e i *\<^sub>R x)" by (rule pdevs_val_zip) also have "\ = (\(i, x)\set ?zip. e i *\<^sub>R x)" by (simp add: sum_list_distinct_conv_sum_set distinct_zipI1) also have [simp]: "set (fst part) \ set (snd part) = {}" by (auto simp: part_def) from partition_set[of "p o snd" ?zip "fst part" "snd part"] have "set ?zip = set (fst part) \ set (snd part)" by (auto simp: part_def) also have "(\a\set (fst part) \ set (snd part). case a of (i, x) \ e i *\<^sub>R x) = (\(i, x)\set (fst part). e i *\<^sub>R x) + (\(i, x)\set (snd part). e i *\<^sub>R x)" by (auto simp: split_beta sum_Un) also have "(\(i, x)\set (fst part). e i *\<^sub>R x) = (\(i, x)\(fst part). e i *\<^sub>R x)" by (simp add: sum_list_distinct_conv_sum_set distinct_zipI1 part_def) also have "\ = (\i e i *\<^sub>R x)" by (subst sum_list_sum_nth) (simp add: split_beta' atLeast0LessThan) also have "\ = pdevs_val (\n. e (map fst (fst part) ! n)) (pdevs_of_list (map snd (fst part)))" by (force simp: pdevs_val_zip sum_list_distinct_conv_sum_set distinct_zipI1 split_beta' in_set_zip intro!: sum.reindex_cong[where l=fst] image_eqI[where x = "(x, map snd (fst part) ! x)" for x]) also have "(\(i, x)\set (snd part). e i *\<^sub>R x) = (\(i, x)\(snd part). e i *\<^sub>R x)" by (simp add: sum_list_distinct_conv_sum_set distinct_zipI1 part_def) also have "\ = (\i e i *\<^sub>R x)" by (subst sum_list_sum_nth) (simp add: split_beta' atLeast0LessThan) also have "\ = pdevs_val (\n. e (map fst (snd part) ! n)) (pdevs_of_list (map snd (snd part)))" by (force simp: pdevs_val_zip sum_list_distinct_conv_sum_set distinct_zipI1 split_beta' in_set_zip intro!: sum.reindex_cong[where l=fst] image_eqI[where x = "(x, map snd (snd part) ! x)" for x]) also have "pdevs_val (\n. e (map fst (fst part) ! n)) (pdevs_of_list (map snd (fst part))) = pdevs_val (\n. if n < degree (pdevs_of_list (map snd (fst part))) then e (map fst (fst part) ! n) else i) (pdevs_of_list (map snd (fst part)))" by (rule pdevs_val_degree_cong) simp_all also have "pdevs_val (\n. e (map fst (snd part) ! n)) (pdevs_of_list (map snd (snd part))) = pdevs_val (\n. if n < degree (pdevs_of_list (map snd (snd part))) then e (map fst (snd part) ! n) else i) (pdevs_of_list (map snd (snd part)))" by (rule pdevs_val_degree_cong) simp_all also have "map snd (snd part) = filter (Not o p) xs" by (simp add: part_def filter_map[symmetric] o_assoc) also have "map snd (fst part) = filter p xs" by (simp add: part_def filter_map[symmetric]) finally show "pdevs_val e (pdevs_of_list xs) = pdevs_val ?f (pdevs_of_list (filter p xs)) + pdevs_val ?g (pdevs_of_list (filter (Not \ p) xs))" . show "?f \ UNIV \ I" "?g \ UNIV \ I" using assms \i\I\ by (auto simp: Pi_iff) qed qed lemma pdevs_apply_pdevs_of_list_append: "pdevs_apply (pdevs_of_list (xs @ zs)) i = (if i < length xs then pdevs_apply (pdevs_of_list xs) i else pdevs_apply (pdevs_of_list zs) (i - length xs))" by (auto simp: pdevs_apply_pdevs_of_list nth_append) lemma degree_pdevs_of_list_le_length[intro, simp]: "degree (pdevs_of_list xs) \ length xs" by (metis less_irrefl_nat le_less_linear less_degree_pdevs_of_list_imp_less_length) lemma degree_pdevs_of_list_append: "degree (pdevs_of_list (xs @ ys)) \ length xs + degree (pdevs_of_list ys)" by (rule degree_le) (auto simp: pdevs_apply_pdevs_of_list_append) lemma pdevs_val_pdevs_of_list_append: assumes "f \ UNIV \ I" assumes "g \ UNIV \ I" obtains e where "pdevs_val f (pdevs_of_list xs) + pdevs_val g (pdevs_of_list ys) = pdevs_val e (pdevs_of_list (xs @ ys))" "e \ UNIV \ I" proof let ?e = "(\i. if i < length xs then f i else g (i - length xs))" have f: "pdevs_val f (pdevs_of_list xs) = (\i\{..R pdevs_apply (pdevs_of_list (xs @ ys)) i)" by (auto simp: pdevs_val_sum degree_gt pdevs_apply_pdevs_of_list_append intro: sum.mono_neutral_cong_left) have g: "pdevs_val g (pdevs_of_list ys) = (\i=length xs ..R pdevs_apply (pdevs_of_list (xs @ ys)) i)" (is "_ = ?sg") by (auto simp: pdevs_val_sum pdevs_apply_pdevs_of_list_append intro!: inj_onI image_eqI[where x="length xs + x" for x] sum.reindex_cong[where l="\i. i - length xs"]) show "pdevs_val f (pdevs_of_list xs) + pdevs_val g (pdevs_of_list ys) = pdevs_val ?e (pdevs_of_list (xs @ ys))" unfolding f g by (subst sum.union_disjoint[symmetric]) (force simp: pdevs_val_sum ivl_disj_un degree_pdevs_of_list_append intro!: sum.mono_neutral_cong_right split: if_split_asm)+ show "?e \ UNIV \ I" using assms by (auto simp: Pi_iff) qed lemma sum_general_mono: fixes f::"'a\('b::ordered_ab_group_add)" assumes [simp,intro]: "finite s" "finite t" assumes f: "\x. x \ s - t \ f x \ 0" assumes g: "\x. x \ t - s \ g x \ 0" assumes fg: "\x. x \ s \ t \ f x \ g x" shows "(\x \ s. f x) \ (\x \ t. g x)" proof - have "s = (s - t) \ (s \ t)" and [intro, simp]: "(s - t) \ (s \ t) = {}" by auto hence "(\x \ s. f x) = (\x \ s - t \ s \ t. f x)" using assms by simp also have "\ = (\x \ s - t. f x) + (\x \ s \ t. f x)" by (simp add: sum_Un) also have "(\x \ s - t. f x) \ 0" by (auto intro!: sum_nonpos f) also have "0 \ (\x \ t - s. g x)" by (auto intro!: sum_nonneg g) also have "(\x \ s \ t. f x) \ (\x \ s \ t. g x)" by (auto intro!: sum_mono fg) also have [intro, simp]: "(t - s) \ (s \ t) = {}" by auto hence "sum g (t - s) + sum g (s \ t) = sum g ((t - s) \ (s \ t))" by (simp add: sum_Un) also have "\ = sum g t" by (auto intro!: sum.cong) finally show ?thesis by simp qed lemma pdevs_val_perm_ex: assumes "xs <~~> ys" assumes mem: "e \ UNIV \ I" shows "\e'. e' \ UNIV \ I \ pdevs_val e (pdevs_of_list xs) = pdevs_val e' (pdevs_of_list ys)" using assms proof (induct arbitrary: e) case Nil thus ?case by auto next case (Cons xs ys z) hence "(e \ (+) (Suc 0)) \ UNIV \ I" by auto from Cons(2)[OF this] obtain e' where "e' \ UNIV \ I" "pdevs_val (e \ (+) (Suc 0)) (pdevs_of_list xs) = pdevs_val e' (pdevs_of_list ys)" by metis thus ?case using Cons by (auto intro!: exI[where x="\x. if x = 0 then e 0 else e' (x - 1)"] simp: o_def Pi_iff) next case (trans xs ys zs) thus ?case by metis next case (swap y x l) thus ?case by (auto intro!: exI[where x="\i. if i = 0 then e 1 else if i = 1 then e 0 else e i"] simp: o_def Pi_iff) qed lemma pdevs_val_perm: assumes "xs <~~> ys" assumes mem: "e \ UNIV \ I" obtains e' where "e' \ UNIV \ I" "pdevs_val e (pdevs_of_list xs) = pdevs_val e' (pdevs_of_list ys)" using assms by (metis pdevs_val_perm_ex) lemma set_distinct_permI: "set xs = set ys \ distinct xs \ distinct ys \ xs <~~> ys" by (metis eq_set_perm_remdups remdups_id_iff_distinct) lemmas pdevs_val_permute = pdevs_val_perm[OF set_distinct_permI] lemma partition_permI: "filter p xs @ filter (Not o p) xs <~~> xs" proof (induct xs) case (Cons x xs) have swap_app_Cons: "filter p xs @ x # [a\xs . \ p a] <~~> x # filter p xs @ [a\xs . \ p a]" by (metis perm_sym perm_append_Cons) also have "\ <~~> x#xs" using Cons by auto finally (trans) show ?case using Cons by simp qed simp lemma pdevs_val_eqI: assumes "\i. i \ pdevs_domain y \ i \ pdevs_domain x \ e i *\<^sub>R pdevs_apply x i = f i *\<^sub>R pdevs_apply y i" assumes "\i. i \ pdevs_domain y \ i \ pdevs_domain x \ f i *\<^sub>R pdevs_apply y i = 0" assumes "\i. i \ pdevs_domain x \ i \ pdevs_domain y \ e i *\<^sub>R pdevs_apply x i = 0" shows "pdevs_val e x = pdevs_val f y" using assms by (force simp: pdevs_val_pdevs_domain intro!: sum.reindex_bij_witness_not_neutral[where i=id and j = id and S'="pdevs_domain x - pdevs_domain y" and T'="pdevs_domain y - pdevs_domain x"]) definition filter_pdevs_raw::"(nat \ 'a \ bool) \ (nat \ 'a::real_vector) \ (nat \ 'a)" where "filter_pdevs_raw I X = (\i. if I i (X i) then X i else 0)" lemma filter_pdevs_raw_nonzeros: "{i. filter_pdevs_raw s f i \ 0} = {i. f i \ 0} \ {x. s x (f x)}" by (auto simp: filter_pdevs_raw_def) lift_definition filter_pdevs::"(nat \ 'a \ bool) \ 'a::real_vector pdevs \ 'a pdevs" is filter_pdevs_raw by (simp add: filter_pdevs_raw_nonzeros) lemma pdevs_apply_filter_pdevs[simp]: "pdevs_apply (filter_pdevs I x) i = (if I i (pdevs_apply x i) then pdevs_apply x i else 0)" by transfer (auto simp: filter_pdevs_raw_def) lemma degree_filter_pdevs_le: "degree (filter_pdevs I x) \ degree x" by (rule degree_leI) (simp split: if_split_asm) lemma pdevs_val_filter_pdevs: "pdevs_val e (filter_pdevs I x) = (\i \ {.. {i. I i (pdevs_apply x i)}. e i *\<^sub>R pdevs_apply x i)" by (auto simp: pdevs_val_sum if_distrib sum.inter_restrict degree_filter_pdevs_le degree_gt intro!: sum.mono_neutral_cong_left split: if_split_asm) lemma pdevs_val_filter_pdevs_dom: "pdevs_val e (filter_pdevs I x) = (\i \ pdevs_domain x \ {i. I i (pdevs_apply x i)}. e i *\<^sub>R pdevs_apply x i)" by (auto simp: pdevs_val_pdevs_domain if_distrib sum.inter_restrict degree_filter_pdevs_le degree_gt intro!: sum.mono_neutral_cong_left split: if_split_asm) lemma pdevs_val_filter_pdevs_eval: "pdevs_val e (filter_pdevs p x) = pdevs_val (\i. if p i (pdevs_apply x i) then e i else 0) x" by (auto split: if_split_asm intro!: pdevs_val_eqI) definition "pdevs_applys X i = map (\x. pdevs_apply x i) X" definition "pdevs_vals e X = map (pdevs_val e) X" definition "aform_vals e X = map (aform_val e) X" definition "filter_pdevs_list I X = map (filter_pdevs (\i _. I i (pdevs_applys X i))) X" lemma pdevs_applys_filter_pdevs_list[simp]: "pdevs_applys (filter_pdevs_list I X) i = (if I i (pdevs_applys X i) then pdevs_applys X i else map (\_. 0) X)" by (auto simp: filter_pdevs_list_def o_def pdevs_applys_def) definition "degrees X = Max (insert 0 (degree ` set X))" abbreviation "degree_aforms X \ degrees (map snd X)" lemma degrees_leI: assumes "\x. x \ set X \ degree x \ K" shows "degrees X \ K" using assms by (auto simp: degrees_def intro!: Max.boundedI) lemma degrees_leD: assumes "degrees X \ K" shows "\x. x \ set X \ degree x \ K" using assms by (auto simp: degrees_def intro!: Max.boundedI) lemma degree_filter_pdevs_list_le: "degrees (filter_pdevs_list I x) \ degrees x" by (rule degrees_leI) (auto simp: filter_pdevs_list_def intro!: degree_le dest!: degrees_leD) definition "dense_list_of_pdevs x = map (\i. pdevs_apply x i) [0..(reverse) ordered coefficients as list\ definition "list_of_pdevs x = map (\i. (i, pdevs_apply x i)) (rev (sorted_list_of_set (pdevs_domain x)))" lemma list_of_pdevs_zero_pdevs[simp]: "list_of_pdevs zero_pdevs = []" by (auto simp: list_of_pdevs_def) lemma sum_list_list_of_pdevs: "sum_list (map snd (list_of_pdevs x)) = sum_list (dense_list_of_pdevs x)" by (auto intro!: sum.mono_neutral_cong_left simp add: degree_gt sum_list_distinct_conv_sum_set dense_list_of_pdevs_def list_of_pdevs_def) lemma sum_list_filter_dense_list_of_pdevs[symmetric]: "sum_list (map snd (filter (p o snd) (list_of_pdevs x))) = sum_list (filter p (dense_list_of_pdevs x))" by (auto intro!: sum.mono_neutral_cong_left simp add: degree_gt sum_list_distinct_conv_sum_set dense_list_of_pdevs_def list_of_pdevs_def o_def filter_map) lemma pdevs_of_list_dense_list_of_pdevs: "pdevs_of_list (dense_list_of_pdevs x) = x" by (auto simp: pdevs_apply_pdevs_of_list dense_list_of_pdevs_def pdevs_eqI) lemma pdevs_val_sum_list: "pdevs_val (\_. c) X = c *\<^sub>R sum_list (map snd (list_of_pdevs X))" by (auto simp: pdevs_val_sum sum_list_list_of_pdevs pdevs_val_const_pdevs_of_list[symmetric] pdevs_of_list_dense_list_of_pdevs) lemma list_of_pdevs_all_nonzero: "list_all (\x. x \ 0) (map snd (list_of_pdevs xs))" by (auto simp: list_of_pdevs_def list_all_iff) lemma list_of_pdevs_nonzero: "x \ set (map snd (list_of_pdevs xs)) \ x \ 0" by (auto simp: list_of_pdevs_def) lemma pdevs_of_list_scaleR_0[simp]: fixes xs::"'a::real_vector list" shows "pdevs_of_list (map ((*\<^sub>R) 0) xs) = zero_pdevs" by (auto simp: pdevs_apply_pdevs_of_list intro!: pdevs_eqI) lemma degree_pdevs_of_list_scaleR: "degree (pdevs_of_list (map ((*\<^sub>R) c) xs)) = (if c \ 0 then degree (pdevs_of_list xs) else 0)" by (auto simp: pdevs_apply_pdevs_of_list intro!: degree_cong) lemma list_of_pdevs_eq: "rev (list_of_pdevs X) = (filter ((\) 0 o snd) (map (\i. (i, pdevs_apply X i)) [0..i. if i < d then 1 else 0) (pdevs_of_list xs)" proof - have "sum_list (take d xs) = 1 *\<^sub>R sum_list (take d xs)" by simp also note pdevs_val_const_pdevs_of_list[symmetric] also have "pdevs_val (\_. 1) (pdevs_of_list (take d xs)) = pdevs_val (\i. if i < d then 1 else 0) (pdevs_of_list xs)" by (auto simp: pdevs_apply_pdevs_of_list split: if_split_asm intro!: pdevs_val_eqI) finally show ?thesis . qed lemma zero_in_range_pdevs_apply[intro, simp]: fixes X::"'a::real_vector pdevs" shows "0 \ range (pdevs_apply X)" by (metis degree_gt less_irrefl rangeI) lemma dense_list_in_range: "x \ set (dense_list_of_pdevs X) \ x \ range (pdevs_apply X)" by (auto simp: dense_list_of_pdevs_def) lemma not_in_dense_list_zeroD: assumes "pdevs_apply X i \ set (dense_list_of_pdevs X)" shows "pdevs_apply X i = 0" proof (rule ccontr) assume "pdevs_apply X i \ 0" hence "i < degree X" by (rule degree_gt) thus False using assms by (auto simp: dense_list_of_pdevs_def) qed lemma list_all_list_of_pdevsI: assumes "\i. i \ pdevs_domain X \ P (pdevs_apply X i)" shows "list_all (\x. P x) (map snd (list_of_pdevs X))" using assms by (auto simp: list_all_iff list_of_pdevs_def) lemma pdevs_of_list_map_scaleR: "pdevs_of_list (map (scaleR r) xs) = scaleR_pdevs r (pdevs_of_list xs)" by (auto intro!: pdevs_eqI simp: pdevs_apply_pdevs_of_list) lemma map_permI: assumes "xs <~~> ys" shows "map f xs <~~> map f ys" using assms by induct auto lemma rev_perm: "rev xs <~~> ys \ xs <~~> ys" by (metis perm.trans perm_rev rev_rev_ident) lemma list_of_pdevs_perm_filter_nonzero: "map snd (list_of_pdevs X) <~~> (filter ((\) 0) (dense_list_of_pdevs X))" proof - have zip_map: "zip [0..i. (i, pdevs_apply X i)) [0.. filter ((\) 0 o snd) (zip [0.. map snd (filter ((\) 0 \ snd) (zip [0..) 0 \ snd) (zip [0..) 0) (dense_list_of_pdevs X)" using map_filter[of snd "(\) 0" "(zip [0.. UNIV \ I" assumes "0 \ I" obtains e' where "pdevs_val e (pdevs_of_list (filter p xs)) = pdevs_val e' (pdevs_of_list xs)" "e' \ UNIV \ I" unfolding pdevs_val_filter_pdevs_eval proof - have "(\_::nat. 0) \ UNIV \ I" using assms by simp have "pdevs_val e (pdevs_of_list (filter p xs)) = pdevs_val e (pdevs_of_list (filter p xs)) + pdevs_val (\_. 0) (pdevs_of_list (filter (Not o p) xs))" by (simp add: pdevs_val_sum) also from pdevs_val_pdevs_of_list_append[OF \e \ _\ \(\_. 0) \ _\] obtain e' where "e' \ UNIV \ I" "\ = pdevs_val e' (pdevs_of_list (filter p xs @ filter (Not o p) xs))" by metis note this(2) also from pdevs_val_perm[OF partition_permI \e' \ _\] obtain e'' where "\ = pdevs_val e'' (pdevs_of_list xs)" "e'' \ UNIV \ I" by metis note this(1) finally show ?thesis using \e'' \ _\ .. qed lemma pdevs_val_of_list_of_pdevs: assumes "e \ UNIV \ I" assumes "0 \ I" obtains e' where "pdevs_val e (pdevs_of_list (map snd (list_of_pdevs X))) = pdevs_val e' X" "e' \ UNIV \ I" proof - obtain e' where "e' \ UNIV \ I" and "pdevs_val e (pdevs_of_list (map snd (list_of_pdevs X))) = pdevs_val e' (pdevs_of_list (filter ((\) 0) (dense_list_of_pdevs X)))" by (rule pdevs_val_perm[OF list_of_pdevs_perm_filter_nonzero assms(1)]) note this(2) also from pdevs_val_filter[OF \e' \ _\ \0 \ I\, of "(\) 0" "dense_list_of_pdevs X"] obtain e'' where "e'' \ UNIV \ I" and "\ = pdevs_val e'' (pdevs_of_list (dense_list_of_pdevs X))" by metis note this(2) also have "\ = pdevs_val e'' X" by (simp add: pdevs_of_list_dense_list_of_pdevs) finally show ?thesis using \e'' \ UNIV \ I\ .. qed lemma pdevs_val_of_list_of_pdevs2: assumes "e \ UNIV \ I" obtains e' where "pdevs_val e X = pdevs_val e' (pdevs_of_list (map snd (list_of_pdevs X)))" "e' \ UNIV \ I" proof - from list_of_pdevs_perm_filter_nonzero[of X] have perm: "(filter ((\) 0) (dense_list_of_pdevs X)) <~~> map snd (list_of_pdevs X)" by (simp add: perm_sym) have "pdevs_val e X = pdevs_val e (pdevs_of_list (dense_list_of_pdevs X))" by (simp add: pdevs_of_list_dense_list_of_pdevs) also from pdevs_val_partition[OF \e \ _\, of "dense_list_of_pdevs X" "(\) 0"] obtain f g where "f \ UNIV \ I" "g \ UNIV \ I" "\ = pdevs_val f (pdevs_of_list (filter ((\) 0) (dense_list_of_pdevs X))) + pdevs_val g (pdevs_of_list (filter (Not \ (\) 0) (dense_list_of_pdevs X)))" (is "_ = ?f + ?g") by metis note this(3) also have "pdevs_of_list [x\dense_list_of_pdevs X . x = 0] = zero_pdevs" by (auto intro!: pdevs_eqI simp: pdevs_apply_pdevs_of_list dest!: nth_mem) hence "?g = 0" by (auto simp: o_def ) also obtain e' where "e' \ UNIV \ I" and "?f = pdevs_val e' (pdevs_of_list (map snd (list_of_pdevs X)))" by (rule pdevs_val_perm[OF perm \f \ _\]) note this(2) finally show ?thesis using \e' \ UNIV \ I\ by (auto intro!: that) qed lemma dense_list_of_pdevs_scaleR: "r \ 0 \ map ((*\<^sub>R) r) (dense_list_of_pdevs x) = dense_list_of_pdevs (scaleR_pdevs r x)" by (auto simp: dense_list_of_pdevs_def) lemma degree_pdevs_of_list_eq: "(\x. x \ set xs \ x \ 0) \ degree (pdevs_of_list xs) = length xs" by (cases xs) (auto simp add: pdevs_apply_pdevs_of_list nth_Cons intro!: degree_eqI split: nat.split) lemma dense_list_of_pdevs_pdevs_of_list: "(\x. x \ set xs \ x \ 0) \ dense_list_of_pdevs (pdevs_of_list xs) = xs" by (auto simp: dense_list_of_pdevs_def degree_pdevs_of_list_eq pdevs_apply_pdevs_of_list intro!: nth_equalityI) lemma pdevs_of_list_sum: assumes "distinct xs" assumes "e \ UNIV \ I" obtains f where "f \ UNIV \ I" "pdevs_val e (pdevs_of_list xs) = (\P\set xs. f P *\<^sub>R P)" proof - define f where "f X = e (the (map_of (zip xs [0.. UNIV \ I" by (auto simp: f_def) moreover have "pdevs_val e (pdevs_of_list xs) = (\P\set xs. f P *\<^sub>R P)" by (auto simp add: pdevs_val_zip f_def assms sum_list_distinct_conv_sum_set[symmetric] in_set_zip map_of_zip_upto2_length_eq_nth intro!: sum_list_nth_eqI) ultimately show ?thesis .. qed lemma pdevs_domain_eq_pdevs_of_list: assumes nz: "\x. x \ set (xs) \ x \ 0" shows "pdevs_domain (pdevs_of_list xs) = {0..x. x \ set xs \ x \ 0" shows "length (list_of_pdevs (pdevs_of_list xs)) = length xs" using nz by (auto simp: list_of_pdevs_def pdevs_domain_eq_pdevs_of_list) lemma nth_list_of_pdevs_pdevs_of_list: assumes nz: "\x. x \ set xs \ x \ 0" assumes l: "n < length xs" shows "list_of_pdevs (pdevs_of_list xs) ! n = ((length xs - Suc n), xs ! (length xs - Suc n))" using nz l by (auto simp: list_of_pdevs_def pdevs_domain_eq_pdevs_of_list rev_nth pdevs_apply_pdevs_of_list) lemma list_of_pdevs_pdevs_of_list_eq: "(\x. x \ set xs \ x \ 0) \ list_of_pdevs (pdevs_of_list xs) = zip (rev [0..x. x \ set xs \ x \ 0" shows "sum_list (filter p (map snd (list_of_pdevs (pdevs_of_list xs)))) = sum_list (filter p xs)" using assms by (auto simp: list_of_pdevs_pdevs_of_list_eq rev_filter[symmetric]) lemma sum_list_partition: fixes xs::"'a::comm_monoid_add list" shows "sum_list (filter p xs) + sum_list (filter (Not o p) xs) = sum_list xs" by (induct xs) (auto simp: ac_simps) subsection \2d zonotopes\ definition "prod_of_pdevs x y = binop_pdevs Pair x y" lemma apply_pdevs_prod_of_pdevs[simp]: "pdevs_apply (prod_of_pdevs x y) i = (pdevs_apply x i, pdevs_apply y i)" unfolding prod_of_pdevs_def by (simp add: zero_prod_def) lemma pdevs_domain_prod_of_pdevs[simp]: "pdevs_domain (prod_of_pdevs x y) = pdevs_domain x \ pdevs_domain y" by (auto simp: zero_prod_def) lemma pdevs_val_prod_of_pdevs[simp]: "pdevs_val e (prod_of_pdevs x y) = (pdevs_val e x, pdevs_val e y)" proof - have "pdevs_val e x = (\i\pdevs_domain x \ pdevs_domain y. e i *\<^sub>R pdevs_apply x i)" (is "_ = ?x") unfolding pdevs_val_pdevs_domain by (rule sum.mono_neutral_cong_left) auto moreover have "pdevs_val e y = (\i\pdevs_domain x \ pdevs_domain y. e i *\<^sub>R pdevs_apply y i)" (is "_ = ?y") unfolding pdevs_val_pdevs_domain by (rule sum.mono_neutral_cong_left) auto ultimately have "(pdevs_val e x, pdevs_val e y) = (?x, ?y)" by auto also have "\ = pdevs_val e (prod_of_pdevs x y)" by (simp add: sum_prod pdevs_val_pdevs_domain) finally show ?thesis by simp qed definition prod_of_aforms (infixr "\\<^sub>a" 80) where "prod_of_aforms x y = ((fst x, fst y), prod_of_pdevs (snd x) (snd y))" subsection \Intervals\ definition One_pdevs_raw::"nat \ 'a::executable_euclidean_space" where "One_pdevs_raw i = (if i < length (Basis_list::'a list) then Basis_list ! i else 0)" lemma zeros_One_pdevs_raw: "One_pdevs_raw -` {0::'a::executable_euclidean_space} = {length (Basis_list::'a list)..}" by (auto simp: One_pdevs_raw_def nonzero_Basis split: if_split_asm dest!: nth_mem) lemma nonzeros_One_pdevs_raw: "{i. One_pdevs_raw i \ (0::'a::executable_euclidean_space)} = - {length (Basis_list::'a list)..}" using zeros_One_pdevs_raw by blast lift_definition One_pdevs::"'a::executable_euclidean_space pdevs" is One_pdevs_raw by (auto simp: nonzeros_One_pdevs_raw) lemma pdevs_apply_One_pdevs[simp]: "pdevs_apply One_pdevs i = (if i < length (Basis_list::'a::executable_euclidean_space list) then Basis_list ! i else 0::'a)" by transfer (simp add: One_pdevs_raw_def) lemma Max_Collect_less_nat: "Max {i::nat. i < k} = (if k = 0 then Max {} else k - 1)" by (auto intro!: Max_eqI) lemma degree_One_pdevs[simp]: "degree (One_pdevs::'a pdevs) = length (Basis_list::'a::executable_euclidean_space list)" by (auto simp: degree_eq_Suc_max Basis_list_nth_nonzero Max_Collect_less_nat intro!: Max_eqI DIM_positive) definition inner_scaleR_pdevs::"'a::euclidean_space \ 'a pdevs \ 'a pdevs" where "inner_scaleR_pdevs b x = unop_pdevs (\x. (b \ x) *\<^sub>R x) x" lemma pdevs_apply_inner_scaleR_pdevs[simp]: "pdevs_apply (inner_scaleR_pdevs a x) i = (a \ (pdevs_apply x i)) *\<^sub>R (pdevs_apply x i)" by (simp add: inner_scaleR_pdevs_def) lemma degree_inner_scaleR_pdevs_le: "degree (inner_scaleR_pdevs (l::'a::executable_euclidean_space) One_pdevs) \ degree (One_pdevs::'a pdevs)" by (rule degree_leI) (auto simp: inner_scaleR_pdevs_def One_pdevs_raw_def) definition "pdevs_of_ivl l u = scaleR_pdevs (1/2) (inner_scaleR_pdevs (u - l) One_pdevs)" lemma degree_pdevs_of_ivl_le: "degree (pdevs_of_ivl l u::'a::executable_euclidean_space pdevs) \ DIM('a)" using degree_inner_scaleR_pdevs_le by (simp add: pdevs_of_ivl_def) lemma pdevs_apply_pdevs_of_ivl: defines "B \ Basis_list::'a::executable_euclidean_space list" shows "pdevs_apply (pdevs_of_ivl l u) i = (if i < length B then ((u - l)\(B!i)/2)*\<^sub>R(B!i) else 0)" by (auto simp: pdevs_of_ivl_def B_def) lemma deg_length_less_imp[simp]: "k < degree (pdevs_of_ivl l u::'a::executable_euclidean_space pdevs) \ k < length (Basis_list::'a list)" by (metis (no_types, hide_lams) degree_One_pdevs degree_inner_scaleR_pdevs_le degree_scaleR_pdevs dual_order.strict_trans length_Basis_list_pos nat_neq_iff not_le pdevs_of_ivl_def) lemma tdev_pdevs_of_ivl: "tdev (pdevs_of_ivl l u) = \u - l\ /\<^sub>R 2" proof - have "tdev (pdevs_of_ivl l u) = (\i pdevs_apply (pdevs_of_ivl l u) i\)" by (auto simp: tdev_def) also have "\ = (\i = 0..pdevs_apply (pdevs_of_ivl l u) i\)" using degree_pdevs_of_ivl_le[of l u] by (intro sum.mono_neutral_cong_left) auto also have "\ = (\i = 0..((u - l) \ Basis_list ! i / 2) *\<^sub>R Basis_list ! i\)" by (auto simp: pdevs_apply_pdevs_of_ivl) also have "\ = (\b \ Basis_list. \((u - l) \ b / 2) *\<^sub>R b\)" by (auto simp: sum_list_sum_nth) also have "\ = (\b\Basis. \((u - l) \ b / 2) *\<^sub>R b\)" by (auto simp: sum_list_distinct_conv_sum_set) also have "\ = \u - l\ /\<^sub>R 2" by (subst euclidean_representation[symmetric, of "\u - l\ /\<^sub>R 2"]) (simp add: abs_inner abs_scaleR) finally show ?thesis . qed definition "aform_of_ivl l u = ((l + u)/\<^sub>R2, pdevs_of_ivl l u)" definition "aform_of_point x = aform_of_ivl x x" lemma Elem_affine_of_ivl_le: assumes "e \ UNIV \ {-1 .. 1}" assumes "l \ u" shows "l \ aform_val e (aform_of_ivl l u)" proof - have "l = (1 / 2) *\<^sub>R l + (1 / 2) *\<^sub>R l" by (simp add: scaleR_left_distrib[symmetric]) also have "\ = (l + u)/\<^sub>R2 - tdev (pdevs_of_ivl l u)" by (auto simp: assms tdev_pdevs_of_ivl algebra_simps) also have "\ \ aform_val e (aform_of_ivl l u)" using abs_pdevs_val_le_tdev[OF assms(1), of "pdevs_of_ivl l u"] by (auto simp: aform_val_def aform_of_ivl_def minus_le_iff dest!: abs_le_D2) finally show ?thesis . qed lemma Elem_affine_of_ivl_ge: assumes "e \ UNIV \ {-1 .. 1}" assumes "l \ u" shows "aform_val e (aform_of_ivl l u) \ u" proof - have "aform_val e (aform_of_ivl l u) \ (l + u)/\<^sub>R2 + tdev (pdevs_of_ivl l u)" using abs_pdevs_val_le_tdev[OF assms(1), of "pdevs_of_ivl l u"] by (auto simp: aform_val_def aform_of_ivl_def minus_le_iff dest!: abs_le_D1) also have "\ = (1 / 2) *\<^sub>R u + (1 / 2) *\<^sub>R u" by (auto simp: assms tdev_pdevs_of_ivl algebra_simps) also have "\ = u" by (simp add: scaleR_left_distrib[symmetric]) finally show ?thesis . qed lemma map_of_zip_upto_length_eq_nth: assumes "i < length B" assumes "d = length B" shows "(map_of (zip [0.. {l .. u}" obtains e where "e \ UNIV \ {-1 .. 1}" "k = aform_val e (aform_of_ivl l u)" proof atomize_elim define e where [abs_def]: "e i = (let b = if i R 2) \ b) / (((u - l) /\<^sub>R 2) \ b))" for i let ?B = "Basis_list::'a list" have "k = (1 / 2) *\<^sub>R (l + u) + (\b \ Basis. (if (u - l) \ b = 0 then 0 else ((k - (1 / 2) *\<^sub>R (l + u)) \ b)) *\<^sub>R b)" (is "_ = _ + ?dots") using assms by (force simp add: algebra_simps eucl_le[where 'a='a] intro!: euclidean_eqI[where 'a='a]) also have "?dots = (\b \ Basis. (if (u - l) \ b = 0 then 0 else ((k - (1 / 2) *\<^sub>R (l + u)) \ b) *\<^sub>R b))" by (auto intro!: sum.cong) also have "\ = (\b \ ?B. (if (u - l) \ b = 0 then 0 else ((k - (1 / 2) *\<^sub>R (l + u)) \ b) *\<^sub>R b))" by (auto simp: sum_list_distinct_conv_sum_set) also have "\ = (\i = 0.. ?B ! i = 0 then 0 else ((k - (1 / 2) *\<^sub>R (l + u)) \ ?B ! i) *\<^sub>R ?B ! i))" by (auto simp: sum_list_sum_nth) also have "\ = (\i = 0.. Basis_list ! i = 0 then 0 else ((k - (1 / 2) *\<^sub>R (l + u)) \ Basis_list ! i) *\<^sub>R Basis_list ! i))" using degree_inner_scaleR_pdevs_le[of "u - l"] by (intro sum.mono_neutral_cong_right) (auto dest!: degree) also have "(1 / 2) *\<^sub>R (l + u) + (\i = 0.. Basis_list ! i = 0 then 0 else ((k - (1 / 2) *\<^sub>R (l + u)) \ Basis_list ! i) *\<^sub>R Basis_list ! i)) = aform_val e (aform_of_ivl l u)" using degree_inner_scaleR_pdevs_le[of "u - l"] by (auto simp: aform_val_def aform_of_ivl_def pdevs_of_ivl_def map_of_zip_upto_length_eq_nth e_def Let_def pdevs_val_sum field_simps intro!: sum.cong) finally have "k = aform_val e (aform_of_ivl l u)" . moreover { fix k l u::real assume *: "l \ k" "k \ u" let ?m = "l / 2 + u / 2" have "\k - ?m\ \ \if k \ ?m then ?m - l else u - ?m\" using * by auto also have "\ \ \u / 2 - l / 2\" by (auto simp: abs_real_def) finally have "\k - (l / 2 + u / 2)\ \ \u / 2 - l/2\" . } note midpoint_abs = this have "e \ UNIV \ {- 1..1}" using assms unfolding e_def Let_def by (intro Pi_I divide_atLeastAtMost_1_absI) (auto simp: map_of_zip_upto_length_eq_nth eucl_le[where 'a='a] divide_le_eq_1 not_less inner_Basis algebra_simps intro!: midpoint_abs dest!: nth_mem) ultimately show "\e. e \ UNIV \ {- 1..1} \ k = aform_val e (aform_of_ivl l u)" by blast qed lemma Inf_aform_aform_of_ivl: assumes "l \ u" shows "Inf_aform (aform_of_ivl l u) = l" using assms by (auto simp: Inf_aform_def aform_of_ivl_def tdev_pdevs_of_ivl abs_diff_eq1 algebra_simps) (metis field_sum_of_halves scaleR_add_left scaleR_one) lemma Sup_aform_aform_of_ivl: assumes "l \ u" shows "Sup_aform (aform_of_ivl l u) = u" using assms by (auto simp: Sup_aform_def aform_of_ivl_def tdev_pdevs_of_ivl abs_diff_eq1 algebra_simps) (metis field_sum_of_halves scaleR_add_left scaleR_one) lemma Affine_aform_of_ivl: "a \ b \ Affine (aform_of_ivl a b) = {a .. b}" by (force simp: Affine_def valuate_def intro!: Elem_affine_of_ivl_ge Elem_affine_of_ivl_le elim!: in_ivl_affine_of_ivlE) end diff --git a/thys/Affine_Arithmetic/Counterclockwise.thy b/thys/Affine_Arithmetic/Counterclockwise.thy --- a/thys/Affine_Arithmetic/Counterclockwise.thy +++ b/thys/Affine_Arithmetic/Counterclockwise.thy @@ -1,321 +1,321 @@ section \Counterclockwise\ theory Counterclockwise -imports "HOL-Analysis.Ordered_Euclidean_Space" +imports "HOL-Analysis.Multivariate_Analysis" begin text \\label{sec:counterclockwise}\ subsection \Auxiliary Lemmas\ lemma convex3_alt: fixes x y z::"'a::real_vector" assumes "0 \ a" "0 \ b" "0 \ c" "a + b + c = 1" obtains u v where "a *\<^sub>R x + b *\<^sub>R y + c *\<^sub>R z = x + u *\<^sub>R (y - x) + v *\<^sub>R (z - x)" and "0 \ u" "0 \ v" "u + v \ 1" proof - from convex_hull_3[of x y z] have "a *\<^sub>R x + b *\<^sub>R y + c *\<^sub>R z \ convex hull {x, y, z}" using assms by auto also note convex_hull_3_alt finally obtain u v where "a *\<^sub>R x + b *\<^sub>R y + c *\<^sub>R z = x + u *\<^sub>R (y - x) + v *\<^sub>R (z - x)" and uv: "0 \ u" "0 \ v" "u + v \ 1" by auto thus ?thesis .. qed lemma (in ordered_ab_group_add) add_nonpos_eq_0_iff: assumes x: "0 \ x" and y: "0 \ y" shows "x + y = 0 \ x = 0 \ y = 0" proof - from add_nonneg_eq_0_iff[of "-x" "-y"] assms have "- (x + y) = 0 \ - x = 0 \ - y = 0" by simp also have "(- (x + y) = 0) = (x + y = 0)" unfolding neg_equal_0_iff_equal .. finally show ?thesis by simp qed lemma sum_nonpos_eq_0_iff: fixes f :: "'a \ 'b::ordered_ab_group_add" shows "\finite A; \x\A. f x \ 0\ \ sum f A = 0 \ (\x\A. f x = 0)" by (induct set: finite) (simp_all add: add_nonpos_eq_0_iff sum_nonpos) lemma fold_if_in_set: "fold (\x m. if P x m then x else m) xs x \ set (x#xs)" by (induct xs arbitrary: x) auto subsection \Sort Elements of a List\ locale linorder_list0 = fixes le::"'a \ 'a \ bool" begin definition "min_for a b = (if le a b then a else b)" lemma min_for_in[simp]: "x \ S \ y \ S \ min_for x y \ S" by (auto simp: min_for_def) lemma fold_min_eqI1: "fold min_for ys y \ set ys \ fold min_for ys y = y" using fold_if_in_set[of _ ys y] by (auto simp: min_for_def[abs_def]) function selsort where "selsort [] = []" | "selsort (y#ys) = (let xm = fold min_for ys y; xs' = List.remove1 xm (y#ys) in (xm#selsort xs'))" by pat_completeness auto termination by (relation "Wellfounded.measure length") (auto simp: length_remove1 intro!: fold_min_eqI1 dest!: length_pos_if_in_set) lemma in_set_selsort_eq: "x \ set (selsort xs) \ x \ (set xs)" by (induct rule: selsort.induct) (auto simp: Let_def intro!: fold_min_eqI1) lemma set_selsort[simp]: "set (selsort xs) = set xs" using in_set_selsort_eq by blast lemma length_selsort[simp]: "length (selsort xs) = length xs" proof (induct xs rule: selsort.induct) case (2 x xs) from 2[OF refl refl] show ?case unfolding selsort.simps by (auto simp: Let_def length_remove1 simp del: selsort.simps split: if_split_asm intro!: Suc_pred dest!: fold_min_eqI1) qed simp lemma distinct_selsort[simp]: "distinct (selsort xs) = distinct xs" by (auto intro!: card_distinct dest!: distinct_card) lemma selsort_eq_empty_iff[simp]: "selsort xs = [] \ xs = []" by (cases xs) (auto simp: Let_def) inductive sortedP :: "'a list \ bool" where Nil: "sortedP []" | Cons: "\y\set ys. le x y \ sortedP ys \ sortedP (x # ys)" inductive_cases sortedP_Nil: "sortedP []" and sortedP_Cons: "sortedP (x#xs)" inductive_simps sortedP_Nil_iff: "sortedP Nil" and sortedP_Cons_iff: "sortedP (Cons x xs)" lemma sortedP_append_iff: "sortedP (xs @ ys) = (sortedP xs & sortedP ys & (\x \ set xs. \y \ set ys. le x y))" by (induct xs) (auto intro!: Nil Cons elim!: sortedP_Cons) lemma sortedP_appendI: "sortedP xs \ sortedP ys \ (\x y. x \ set xs \ y \ set ys \ le x y) \ sortedP (xs @ ys)" by (induct xs) (auto intro!: Nil Cons elim!: sortedP_Cons) lemma sorted_nth_less: "sortedP xs \ i < j \ j < length xs \ le (xs ! i) (xs ! j)" by (induct xs arbitrary: i j) (auto simp: nth_Cons split: nat.split elim!: sortedP_Cons) lemma sorted_butlastI[intro, simp]: "sortedP xs \ sortedP (butlast xs)" by (induct xs) (auto simp: elim!: sortedP_Cons intro!: sortedP.Cons dest!: in_set_butlastD) lemma sortedP_right_of_append1: assumes "sortedP (zs@[z])" assumes "y \ set zs" shows "le y z" using assms by (induct zs arbitrary: y z) (auto elim!: sortedP_Cons) lemma sortedP_right_of_last: assumes "sortedP zs" assumes "y \ set zs" "y \ last zs" shows "le y (last zs)" using assms apply (intro sortedP_right_of_append1[of "butlast zs" "last zs" y]) subgoal by (metis append_is_Nil_conv list.distinct(1) snoc_eq_iff_butlast split_list) subgoal by (metis List.insert_def append_butlast_last_id insert_Nil list.distinct(1) rotate1.simps(2) set_ConsD set_rotate1) done lemma selsort_singleton_iff: "selsort xs = [x] \ xs = [x]" by (induct xs) (auto simp: Let_def) lemma hd_last_sorted: assumes "sortedP xs" "length xs > 1" shows "le (hd xs) (last xs)" proof (cases xs) case (Cons y ys) note ys = this thus ?thesis using ys assms by (auto elim!: sortedP_Cons) qed (insert assms, simp) end lemma (in comm_monoid_add) sum_list_distinct_selsort: assumes "distinct xs" shows "sum_list (linorder_list0.selsort le xs) = sum_list xs" using assms apply (simp add: distinct_sum_list_conv_Sum linorder_list0.distinct_selsort) apply (rule sum.cong) subgoal by (simp add: linorder_list0.set_selsort) subgoal by simp done declare linorder_list0.sortedP_Nil_iff[code] linorder_list0.sortedP_Cons_iff[code] linorder_list0.selsort.simps[code] linorder_list0.min_for_def[code] locale linorder_list = linorder_list0 le for le::"'a::ab_group_add \ _" + fixes S assumes order_refl: "a \ S \ le a a" assumes trans': "a \ S \ b \ S \ c \ S \ a \ b \ b \ c \ a \ c \ le a b \ le b c \ le a c" assumes antisym: "a \ S \ b \ S \ le a b \ le b a \ a = b" assumes linear': "a \ S \ b \ S \ a \ b \ le a b \ le b a" begin lemma trans: "a \ S \ b \ S \ c \ S \ le a b \ le b c \ le a c" by (cases "a = b" "b = c" "a = c" rule: bool.exhaust[case_product bool.exhaust[case_product bool.exhaust]]) (auto simp: order_refl intro: trans') lemma linear: "a \ S \ b \ S \ le a b \ le b a" by (cases "a = b") (auto simp: linear' order_refl) lemma min_le1: "w \ S \ y \ S \ le (min_for w y) y" and min_le2: "w \ S \ y \ S \ le (min_for w y) w" using linear by (auto simp: min_for_def refl) lemma fold_min: assumes "set xs \ S" shows "list_all (\y. le (fold min_for (tl xs) (hd xs)) y) xs" proof (cases xs) case (Cons y ys) hence subset: "set (y#ys) \ S" using assms by auto show ?thesis unfolding Cons list.sel using subset proof (induct ys arbitrary: y) case (Cons z zs) hence IH: "\y. y \ S \ list_all (le (fold min_for zs y)) (y # zs)" by simp let ?f = "fold min_for zs (min_for z y)" have "?f \ set ((min_for z y)#zs)" unfolding min_for_def[abs_def] by (rule fold_if_in_set) also have "\ \ S" using Cons.prems by auto finally have "?f \ S" . have "le ?f (min_for z y)" using IH[of "min_for z y"] Cons.prems by auto moreover have "le (min_for z y) y" "le (min_for z y) z" using Cons.prems by (auto intro!: min_le1 min_le2) ultimately have "le ?f y" "le ?f z" using Cons.prems \?f \ S\ by (auto intro!: trans[of ?f "min_for z y"]) thus ?case using IH[of "min_for z y"] using Cons.prems by auto qed (simp add: order_refl) qed simp lemma sortedP_selsort: assumes "set xs \ S" shows "sortedP (selsort xs)" using assms proof (induction xs rule: selsort.induct) case (2 z zs) from this fold_min[of "z#zs"] show ?case by (fastforce simp: list_all_iff Let_def simp del: remove1.simps intro: Cons intro!: 2(1)[OF refl refl] dest!: rev_subsetD[OF _ set_remove1_subset])+ qed (auto intro!: Nil) end subsection \Abstract CCW Systems\ locale ccw_system0 = fixes ccw::"'a \ 'a \ 'a \ bool" and S::"'a set" begin abbreviation "indelta t p q r \ ccw t q r \ ccw p t r \ ccw p q t" abbreviation "insquare p q r s \ ccw p q r \ ccw q r s \ ccw r s p \ ccw s p q" end abbreviation "distinct3 p q r \ \(p = q \ p = r \ q = r)" abbreviation "distinct4 p q r s \ \(p = q \ p = r \ p = s \ \ distinct3 q r s)" abbreviation "distinct5 p q r s t \ \(p = q \ p = r \ p = s \ p = t \ \ distinct4 q r s t)" abbreviation "in3 S p q r \ p \ S \ q \ S \ r \ S" abbreviation "in4 S p q r s \ in3 S p q r \ s \ S" abbreviation "in5 S p q r s t \ in4 S p q r s \ t \ S" locale ccw_system12 = ccw_system0 + assumes cyclic: "ccw p q r \ ccw q r p" assumes ccw_antisym: "distinct3 p q r \ in3 S p q r \ ccw p q r \ \ ccw p r q" locale ccw_system123 = ccw_system12 + assumes nondegenerate: "distinct3 p q r \ in3 S p q r \ ccw p q r \ ccw p r q" begin lemma not_ccw_eq: "distinct3 p q r \ in3 S p q r \ \ ccw p q r \ ccw p r q" using ccw_antisym nondegenerate by blast end locale ccw_system4 = ccw_system123 + assumes interior: "distinct4 p q r t \ in4 S p q r t \ ccw t q r \ ccw p t r \ ccw p q t \ ccw p q r" begin lemma interior': "distinct4 p q r t \ in4 S p q r t \ ccw p q t \ ccw q r t \ ccw r p t \ ccw p q r" by (metis ccw_antisym cyclic interior nondegenerate) end locale ccw_system1235' = ccw_system123 + assumes dual_transitive: "distinct5 p q r s t \ in5 S p q r s t \ ccw s t p \ ccw s t q \ ccw s t r \ ccw t p q \ ccw t q r \ ccw t p r" locale ccw_system1235 = ccw_system123 + assumes transitive: "distinct5 p q r s t \ in5 S p q r s t \ ccw t s p \ ccw t s q \ ccw t s r \ ccw t p q \ ccw t q r \ ccw t p r" begin lemmas ccw_axioms = cyclic nondegenerate ccw_antisym transitive sublocale ccw_system1235' proof (unfold_locales, rule ccontr, goal_cases) case prems: (1 p q r s t) hence "ccw s p q \ ccw s r p" by (metis ccw_axioms prems) moreover have "ccw s r p \ ccw s q r" by (metis ccw_axioms prems) moreover have "ccw s q r \ ccw s p q" by (metis ccw_axioms prems) ultimately have "ccw s p q \ ccw s r p \ ccw s q r \ ccw s q p \ ccw s p r \ ccw s r q" by (metis ccw_axioms prems) thus False by (metis ccw_axioms prems) qed end locale ccw_system = ccw_system1235 + ccw_system4 end diff --git a/thys/Affine_Arithmetic/Executable_Euclidean_Space.thy b/thys/Affine_Arithmetic/Executable_Euclidean_Space.thy --- a/thys/Affine_Arithmetic/Executable_Euclidean_Space.thy +++ b/thys/Affine_Arithmetic/Executable_Euclidean_Space.thy @@ -1,1077 +1,1077 @@ section \Euclidean Space: Executability\ theory Executable_Euclidean_Space imports - "HOL-Analysis.Ordered_Euclidean_Space" + "HOL-Analysis.Multivariate_Analysis" "List-Index.List_Index" "HOL-Library.Float" Affine_Arithmetic_Auxiliarities begin subsection \Ordered representation of Basis and Rounding of Components\ class executable_euclidean_space = ordered_euclidean_space + fixes Basis_list eucl_down eucl_truncate_down eucl_truncate_up assumes eucl_down_def: "eucl_down p b = (\i \ Basis. round_down p (b \ i) *\<^sub>R i)" assumes eucl_truncate_down_def: "eucl_truncate_down q b = (\i \ Basis. truncate_down q (b \ i) *\<^sub>R i)" assumes eucl_truncate_up_def: "eucl_truncate_up q b = (\i \ Basis. truncate_up q (b \ i) *\<^sub>R i)" assumes Basis_list[simp]: "set Basis_list = Basis" assumes distinct_Basis_list[simp]: "distinct Basis_list" begin lemma length_Basis_list: "length Basis_list = card Basis" by (metis Basis_list distinct_Basis_list distinct_card) end lemma eucl_truncate_down_zero[simp]: "eucl_truncate_down p 0 = 0" by (auto simp: eucl_truncate_down_def truncate_down_zero) lemma eucl_truncate_up_zero[simp]: "eucl_truncate_up p 0 = 0" by (auto simp: eucl_truncate_up_def) subsection \Instantiations\ instantiation real::executable_euclidean_space begin definition Basis_list_real :: "real list" where "Basis_list_real = [1]" definition "eucl_down prec b = round_down prec b" definition "eucl_truncate_down prec b = truncate_down prec b" definition "eucl_truncate_up prec b = truncate_up prec b" instance proof qed (auto simp: Basis_list_real_def eucl_down_real_def eucl_truncate_down_real_def eucl_truncate_up_real_def) end instantiation prod::(executable_euclidean_space, executable_euclidean_space) executable_euclidean_space begin definition Basis_list_prod :: "('a \ 'b) list" where "Basis_list_prod = zip Basis_list (replicate (length (Basis_list::'a list)) 0) @ zip (replicate (length (Basis_list::'b list)) 0) Basis_list" definition "eucl_down p a = (eucl_down p (fst a), eucl_down p (snd a))" definition "eucl_truncate_down p a = (eucl_truncate_down p (fst a), eucl_truncate_down p (snd a))" definition "eucl_truncate_up p a = (eucl_truncate_up p (fst a), eucl_truncate_up p (snd a))" instance proof show "set Basis_list = (Basis::('a*'b) set)" by (auto simp: Basis_list_prod_def Basis_prod_def elim!: in_set_zipE) (auto simp: Basis_list[symmetric] in_set_zip in_set_conv_nth simp del: Basis_list) show "distinct (Basis_list::('a*'b)list)" using distinct_Basis_list[where 'a='a] distinct_Basis_list[where 'a='b] by (auto simp: Basis_list_prod_def Basis_list intro: distinct_zipI1 distinct_zipI2 elim!: in_set_zipE) qed (auto simp: eucl_down_prod_def eucl_truncate_down_prod_def eucl_truncate_up_prod_def sum_Basis_prod_eq inner_add_left inner_sum_left inner_Basis eucl_down_def eucl_truncate_down_def eucl_truncate_up_def intro!: euclidean_eqI[where 'a="'a*'b"]) end lemma eucl_truncate_down_Basis[simp]: "i \ Basis \ eucl_truncate_down e x \ i = truncate_down e (x \ i)" by (simp add: eucl_truncate_down_def) lemma eucl_truncate_down_correct: "dist (x::'a::executable_euclidean_space) (eucl_down e x) \ {0..sqrt (DIM('a)) * 2 powr of_int (- e)}" proof - have "dist x (eucl_down e x) = sqrt (\i\Basis. (dist (x \ i) (eucl_down e x \ i))\<^sup>2)" unfolding euclidean_dist_l2[where 'a='a] L2_set_def .. also have "\ \ sqrt (\i\(Basis::'a set). ((2 powr of_int (- e))\<^sup>2))" by (intro real_sqrt_le_mono sum_mono power_mono) (auto simp: dist_real_def eucl_down_def abs_round_down_le) finally show ?thesis by (simp add: real_sqrt_mult) qed lemma eucl_down: "eucl_down e (x::'a::executable_euclidean_space) \ x" by (auto simp add: eucl_le[where 'a='a] round_down eucl_down_def) lemma eucl_truncate_down: "eucl_truncate_down e (x::'a::executable_euclidean_space) \ x" by (auto simp add: eucl_le[where 'a='a] truncate_down) lemma eucl_truncate_down_le: "x \ y \ eucl_truncate_down w x \ (y::'a::executable_euclidean_space)" using eucl_truncate_down by (rule order.trans) lemma eucl_truncate_up_Basis[simp]: "i \ Basis \ eucl_truncate_up e x \ i = truncate_up e (x \ i)" by (simp add: eucl_truncate_up_def truncate_up_def) lemma eucl_truncate_up: "x \ eucl_truncate_up e (x::'a::executable_euclidean_space)" by (auto simp add: eucl_le[where 'a='a] round_up truncate_up_def) lemma eucl_truncate_up_le: "x \ y \ x \ eucl_truncate_up e (y::'a::executable_euclidean_space)" using _ eucl_truncate_up by (rule order.trans) lemma eucl_truncate_down_mono: fixes x::"'a::executable_euclidean_space" shows "x \ y \ eucl_truncate_down p x \ eucl_truncate_down p y" by (auto simp: eucl_le[where 'a='a] intro!: truncate_down_mono) lemma eucl_truncate_up_mono: fixes x::"'a::executable_euclidean_space" shows "x \ y \ eucl_truncate_up p x \ eucl_truncate_up p y" by (auto simp: eucl_le[where 'a='a] intro!: truncate_up_mono) lemma infnorm[code]: fixes x::"'a::executable_euclidean_space" shows "infnorm x = fold max (map (\i. abs (x \ i)) Basis_list) 0" by (auto simp: Max.set_eq_fold[symmetric] infnorm_Max[symmetric] infnorm_pos_le intro!: max.absorb2[symmetric]) declare Inf_real_def[code del] declare Sup_real_def[code del] declare Inf_prod_def[code del] declare Sup_prod_def[code del] declare [[code abort: "Inf::real set \ real"]] declare [[code abort: "Sup::real set \ real"]] declare [[code abort: "Inf::('a::Inf * 'b::Inf) set \ 'a * 'b"]] declare [[code abort: "Sup::('a::Sup * 'b::Sup) set \ 'a * 'b"]] lemma nth_Basis_list_in_Basis[simp]: "n < length (Basis_list::'a::executable_euclidean_space list) \ Basis_list ! n \ (Basis::'a set)" by (metis Basis_list nth_mem) subsection \Representation as list\ lemma nth_eq_iff_index: "distinct xs \ n < length xs \ xs ! n = i \ n = index xs i" using index_nth_id by fastforce lemma in_Basis_index_Basis_list: "i \ Basis \ i = Basis_list ! index Basis_list i" by simp lemmas [simp] = length_Basis_list lemma sum_Basis_sum_nth_Basis_list: "(\i\Basis. f i) = (\i(x, i)\zip xs Basis_list. x *\<^sub>R i)" lemma eucl_of_list_nth: assumes "length xs = DIM('a)" shows "eucl_of_list xs = (\iR ((Basis_list::'a list) ! i))" by (auto simp: eucl_of_list_def sum_list_sum_nth length_Basis_list assms atLeast0LessThan) lemma eucl_of_list_inner: fixes i::"'a::executable_euclidean_space" assumes i: "i \ Basis" assumes l: "length xs = DIM('a)" shows "eucl_of_list xs \ i = xs ! (index Basis_list i)" by (simp add: eucl_of_list_nth[OF l] inner_sum_left assms inner_Basis nth_eq_iff_index sum.delta if_distrib cong: if_cong) lemma inner_eucl_of_list: fixes i::"'a::executable_euclidean_space" assumes i: "i \ Basis" assumes l: "length xs = DIM('a)" shows "i \ eucl_of_list xs = xs ! (index Basis_list i)" using eucl_of_list_inner[OF assms] by (auto simp: inner_commute) definition "list_of_eucl x = map ((\) x) Basis_list" lemma index_Basis_list_nth[simp]: "i < DIM('a::executable_euclidean_space) \ index Basis_list ((Basis_list::'a list) ! i) = i" by (simp add: index_nth_id) lemma list_of_eucl_eucl_of_list[simp]: "length xs = DIM('a::executable_euclidean_space) \ list_of_eucl (eucl_of_list xs::'a) = xs" by (auto simp: list_of_eucl_def eucl_of_list_inner intro!: nth_equalityI) lemma eucl_of_list_list_of_eucl[simp]: "eucl_of_list (list_of_eucl x) = x" by (auto simp: list_of_eucl_def eucl_of_list_inner intro!: euclidean_eqI[where 'a='a]) lemma length_list_of_eucl[simp]: "length (list_of_eucl (x::'a::executable_euclidean_space)) = DIM('a)" by (auto simp: list_of_eucl_def) lemma list_of_eucl_nth[simp]: "n < DIM('a::executable_euclidean_space) \ list_of_eucl x ! n = x \ (Basis_list ! n::'a)" by (auto simp: list_of_eucl_def) lemma nth_ge_len: "n \ length xs \ xs ! n = [] ! (n - length xs)" by (induction xs arbitrary: n) auto lemma list_of_eucl_nth_if: "list_of_eucl x ! n = (if n < DIM('a::executable_euclidean_space) then x \ (Basis_list ! n::'a) else [] ! (n - DIM('a)))" apply (auto simp: list_of_eucl_def ) apply (subst nth_ge_len) apply auto done lemma list_of_eucl_eq_iff: "list_of_eucl (x::'a::executable_euclidean_space) = list_of_eucl (y::'b::executable_euclidean_space) \ (DIM('a) = DIM('b) \ (\i < DIM('b). x \ Basis_list ! i = y \ Basis_list ! i))" by (auto simp: list_eq_iff_nth_eq) lemma eucl_le_Basis_list_iff: "(x::'a::executable_euclidean_space) \ y \ (\i Basis_list ! i \ y \ Basis_list ! i)" apply (auto simp: eucl_le[where 'a='a]) subgoal for i subgoal by (auto dest!: spec[where x="index Basis_list i"]) done done lemma eucl_of_list_inj: "length xs = DIM('a::executable_euclidean_space) \ length ys = DIM('a) \ (eucl_of_list xs::'a) = eucl_of_list (ys) \ xs = ys" apply (auto intro!: nth_equalityI simp: euclidean_eq_iff[where 'a="'a"] eucl_of_list_inner) using nth_Basis_list_in_Basis[where 'a="'a"] by fastforce lemma eucl_of_list_map_plus[simp]: assumes [simp]: "length xs = DIM('a::executable_euclidean_space)" shows "(eucl_of_list (map (\x. f x + g x) xs)::'a) = eucl_of_list (map f xs) + eucl_of_list (map g xs)" by (auto intro!: euclidean_eqI[where 'a='a] simp: algebra_simps eucl_of_list_inner) lemma eucl_of_list_map_uminus[simp]: assumes [simp]: "length xs = DIM('a::executable_euclidean_space)" shows "(eucl_of_list (map (\x. - f x) xs)::'a) = - eucl_of_list (map f xs)" by (auto intro!: euclidean_eqI[where 'a='a] simp: algebra_simps eucl_of_list_inner) lemma eucl_of_list_map_mult_left[simp]: assumes [simp]: "length xs = DIM('a::executable_euclidean_space)" shows "(eucl_of_list (map (\x. r * f x) xs)::'a) = r *\<^sub>R eucl_of_list (map f xs)" by (auto intro!: euclidean_eqI[where 'a='a] simp: algebra_simps eucl_of_list_inner) lemma eucl_of_list_map_mult_right[simp]: assumes [simp]: "length xs = DIM('a::executable_euclidean_space)" shows "(eucl_of_list (map (\x. f x * r) xs)::'a) = r *\<^sub>R eucl_of_list (map f xs)" by (auto intro!: euclidean_eqI[where 'a='a] simp: algebra_simps eucl_of_list_inner) lemma eucl_of_list_map_divide_right[simp]: assumes [simp]: "length xs = DIM('a::executable_euclidean_space)" shows "(eucl_of_list (map (\x. f x / r) xs)::'a) = eucl_of_list (map f xs) /\<^sub>R r" by (auto intro!: euclidean_eqI[where 'a='a] simp: algebra_simps eucl_of_list_inner divide_simps) lemma eucl_of_list_map_const[simp]: assumes [simp]: "length xs = DIM('a::executable_euclidean_space)" shows "(eucl_of_list (map (\x. c) xs)::'a) = c *\<^sub>R One" by (auto intro!: euclidean_eqI[where 'a='a] simp: algebra_simps eucl_of_list_inner) lemma replicate_eq_list_of_eucl_zero: "replicate DIM('a::executable_euclidean_space) 0 = list_of_eucl (0::'a)" by (auto intro!: nth_equalityI) lemma eucl_of_list_append_zeroes[simp]: "eucl_of_list (xs @ replicate n 0) = eucl_of_list xs" unfolding eucl_of_list_def apply (auto simp: sum_list_sum_nth) apply (rule sum.mono_neutral_cong_right) by (auto simp: nth_append) lemma Basis_prodD: assumes "(i, j) \ Basis" shows "i \ Basis \ j = 0 \ i = 0 \ j \ Basis" using assms by (auto simp: Basis_prod_def) lemma eucl_of_list_take_DIM[simp]: assumes "d = DIM('b::executable_euclidean_space)" shows "(eucl_of_list (take d xs)::'b) = (eucl_of_list xs)" by (auto simp: eucl_of_list_inner eucl_of_list_def fst_sum_list sum_list_sum_nth assms dest!: Basis_prodD) lemma eucl_of_list_eqI: assumes "take DIM('a) (xs @ replicate (DIM('a) - length xs) 0) = take DIM('a) (ys @ replicate (DIM('a) - length ys) 0)" shows "eucl_of_list xs = (eucl_of_list ys::'a::executable_euclidean_space)" proof - have "(eucl_of_list xs::'a) = eucl_of_list (take DIM('a) (xs @ replicate (DIM('a) - length xs) 0))" by (simp add: ) also note assms also have "eucl_of_list (take DIM('a) (ys @ replicate (DIM('a) - length ys) 0)) = (eucl_of_list ys::'a)" by simp finally show ?thesis . qed lemma eucl_of_list_replicate_zero[simp]: "eucl_of_list (replicate E 0) = 0" proof - have "eucl_of_list (replicate E 0) = (eucl_of_list (replicate E 0 @ replicate (DIM('a) - E) 0)::'a)" by simp also have "\ = eucl_of_list (replicate DIM('a) 0)" apply (rule eucl_of_list_eqI) by (auto simp: min_def nth_append intro!: nth_equalityI) also have "\ = 0" by (simp add: replicate_eq_list_of_eucl_zero) finally show ?thesis by simp qed lemma eucl_of_list_Nil[simp]: "eucl_of_list [] = 0" using eucl_of_list_replicate_zero[of 0] by simp lemma fst_eucl_of_list_prod: shows "fst (eucl_of_list xs::'b::executable_euclidean_space \ _) = (eucl_of_list (take DIM('b) xs)::'b)" apply (auto simp: eucl_of_list_inner eucl_of_list_def fst_sum_list dest!: Basis_prodD) apply (simp add: sum_list_sum_nth) apply (rule sum.mono_neutral_cong_right) subgoal by simp subgoal by auto subgoal by (auto simp: Basis_list_prod_def nth_append) subgoal by (auto simp: Basis_list_prod_def nth_append) done lemma index_zip_replicate1[simp]: "index (zip (replicate d a) bs) (a, b) = index bs b" if "d = length bs" using that by (induction bs arbitrary: d) auto lemma index_zip_replicate2[simp]: "index (zip as (replicate d b)) (a, b) = index as a" if "d = length as" using that by (induction as arbitrary: d) auto lemma index_Basis_list_prod[simp]: fixes a::"'a::executable_euclidean_space" and b::"'b::executable_euclidean_space" shows "a \ Basis \ index Basis_list (a, 0::'b) = index Basis_list a" "b \ Basis \ index Basis_list (0::'a, b) = DIM('a) + index Basis_list b" by (auto simp: Basis_list_prod_def index_append in_set_zip zip_replicate index_map_inj dest: spec[where x="index Basis_list a"]) lemma eucl_of_list_eq_takeI: assumes "(eucl_of_list (take DIM('a::executable_euclidean_space) xs)::'a) = x" shows "eucl_of_list xs = x" using eucl_of_list_take_DIM[OF refl, of xs, where 'b='a] assms by auto lemma eucl_of_list_inner_le: fixes i::"'a::executable_euclidean_space" assumes i: "i \ Basis" assumes l: "length xs \ DIM('a)" shows "eucl_of_list xs \ i = xs ! (index Basis_list i)" proof - have "(eucl_of_list xs::'a) = eucl_of_list (take DIM('a) (xs @ (replicate (DIM('a) - length xs) 0)))" by (rule eucl_of_list_eq_takeI) simp also have "\ \ i = xs ! (index Basis_list i)" using assms by (subst eucl_of_list_inner) auto finally show ?thesis . qed lemma eucl_of_list_prod_if: assumes "length xs = DIM('a::executable_euclidean_space) + DIM('b::executable_euclidean_space)" shows "eucl_of_list xs = (eucl_of_list (take DIM('a) xs)::'a, eucl_of_list (drop DIM('a) xs)::'b)" apply (rule euclidean_eqI) using assms apply (auto simp: eucl_of_list_inner dest!: Basis_prodD) apply (subst eucl_of_list_inner_le) apply (auto simp: Basis_list_prod_def index_append in_set_zip) done lemma snd_eucl_of_list_prod: shows "snd (eucl_of_list xs::'b::executable_euclidean_space \ 'c::executable_euclidean_space) = (eucl_of_list (drop DIM('b) xs)::'c)" proof (cases "length xs \ DIM('b)") case True then show ?thesis by (auto simp: eucl_of_list_inner eucl_of_list_def snd_sum_list dest!: Basis_prodD) (simp add: sum_list_sum_nth Basis_list_prod_def nth_append) next case False have "xs = take DIM('b) xs @ drop DIM('b) xs" by simp also have "eucl_of_list \ = (eucl_of_list (\ @ replicate (length xs - DIM('c)) 0)::'b \ 'c)" by simp finally have "eucl_of_list xs = (eucl_of_list (xs @ replicate (DIM('b) + DIM('c) - length xs) 0)::'b \ 'c)" by simp also have "\ = eucl_of_list (take (DIM ('b \ 'c)) (xs @ replicate (DIM('b) + DIM('c) - length xs) 0))" by (simp add: ) finally have *: "(eucl_of_list xs::'b\'c) = eucl_of_list (take DIM('b \ 'c) (xs @ replicate (DIM('b) + DIM('c) - length xs) 0))" by simp show ?thesis apply (subst *) apply (subst eucl_of_list_prod_if) subgoal by simp subgoal apply simp apply (subst (2) eucl_of_list_take_DIM[OF refl, symmetric]) apply (subst (2) eucl_of_list_take_DIM[OF refl, symmetric]) apply (rule arg_cong[where f=eucl_of_list]) by (auto intro!: nth_equalityI simp: nth_append min_def split: if_splits) done qed lemma eucl_of_list_prod: shows "eucl_of_list xs = (eucl_of_list (take DIM('b) xs)::'b::executable_euclidean_space, eucl_of_list (drop DIM('b) xs)::'c::executable_euclidean_space)" using snd_eucl_of_list_prod[of xs, where 'b='b and 'c='c] using fst_eucl_of_list_prod[of xs, where 'b='b and 'a='c] by (auto simp del: snd_eucl_of_list_prod fst_eucl_of_list_prod simp add: prod_eq_iff) lemma eucl_of_list_real[simp]: "eucl_of_list [x] = (x::real)" by (auto simp: eucl_of_list_def Basis_list_real_def) lemma eucl_of_list_append[simp]: assumes "length xs = DIM('i::executable_euclidean_space)" assumes "length ys = DIM('j::executable_euclidean_space)" shows "eucl_of_list (xs @ ys) = (eucl_of_list xs::'i, eucl_of_list ys::'j)" using assms by (auto simp: eucl_of_list_prod) lemma list_allI: "(\x. x \ set xs \ P x) \ list_all P xs" by (auto simp: list_all_iff) lemma concat_map_nthI: assumes "\x y. x \ set xs \ y \ set (f x) \ P y" assumes "j < length (concat (map f xs))" shows "P (concat (map f xs) ! j)" proof - have "list_all P (concat (map f xs))" by (rule list_allI) (auto simp: assms) then show ?thesis by (auto simp: list_all_length assms) qed lemma map_nth_append1: assumes "length xs = d" shows "map ((!) (xs @ ys)) [0.. y = z" by (metis eucl_of_list_list_of_eucl) lemma length_Basis_list_pos[simp]: "length Basis_list > 0" by (metis length_pos_if_in_set Basis_list SOME_Basis) lemma Basis_list_nth_nonzero: "i < length (Basis_list::'a::executable_euclidean_space list) \ (Basis_list::'a list) ! i \ 0" by (auto dest!: nth_mem simp: nonzero_Basis) lemma nth_Basis_list_prod: "i < DIM('a) + DIM('b) \ (Basis_list::('a::executable_euclidean_space \ 'b::executable_euclidean_space) list) ! i = (if i < DIM('a) then (Basis_list ! i, 0) else (0, Basis_list ! (i - DIM('a))))" by (auto simp: Basis_list_nth_nonzero prod_eq_iff Basis_list_prod_def nth_append not_less) lemma eucl_of_list_if: assumes [simp]: "length xs = DIM('a::executable_euclidean_space)" "distinct xs" shows "eucl_of_list (map (\xa. if xa = x then 1 else 0) (xs::nat list)) = (if x \ set xs then Basis_list ! index xs x else 0::'a)" by (rule euclidean_eqI) (auto simp: eucl_of_list_inner inner_Basis index_nth_id) lemma take_append_take_minus_idem: "take n XS @ map ((!) XS) [n..b\Basis. f b)" by (subst sum_list_distinct_conv_sum_set) (auto simp: Basis_list distinct_Basis_list) lemma hd_Basis_list[simp]: "hd Basis_list \ Basis" unfolding Basis_list[symmetric] by (rule hd_in_set) (auto simp: set_empty[symmetric]) definition "inner_lv_rel a b = sum_list (map2 (*) a b)" lemma eucl_of_list_inner_eq: "(eucl_of_list xs::'a) \ eucl_of_list ys = inner_lv_rel xs ys" if "length xs = DIM('a::executable_euclidean_space)" "length ys = DIM('a)" using that by (subst euclidean_inner[abs_def], subst sum_list_Basis_list[symmetric]) (auto simp: eucl_of_list_inner sum_list_sum_nth index_nth_id inner_lv_rel_def) lemma euclidean_vec_componentwise: "(\(xa::'a::euclidean_space^'b::finite)\Basis. f xa) = (\a\Basis. (\b::'b\UNIV. f (axis b a)))" apply (auto simp: Basis_vec_def) apply (subst sum.swap) apply (subst sum.Union_disjoint) apply auto apply (simp add: axis_eq_axis nonzero_Basis) apply (simp add: axis_eq_axis nonzero_Basis) apply (subst sum.reindex) apply (auto intro!: injI) subgoal apply (auto simp: set_eq_iff) by (metis (full_types) all_not_in_conv inner_axis_axis inner_eq_zero_iff nonempty_Basis nonzero_Basis) apply (rule sum.cong[OF refl]) apply (auto ) apply (rule sum.reindex_cong[OF _ _ refl]) apply (auto intro!: inj_onI) using axis_eq_axis by blast lemma vec_nth_inner_scaleR_craziness: "f (x $ i \ j) *\<^sub>R j = (\xa\UNIV. f (x $ xa \ j) *\<^sub>R axis xa j) $ i" by vector (auto simp: axis_def if_distrib scaleR_vec_def sum.delta' cong: if_cong) instantiation vec :: ("{executable_euclidean_space}", enum) executable_euclidean_space begin definition Basis_list_vec :: "('a, 'b) vec list" where "Basis_list_vec = concat (map (\n. map (axis n) Basis_list) enum_class.enum)" definition eucl_down_vec :: "int \ ('a, 'b) vec \ ('a, 'b) vec" where "eucl_down_vec p x = (\ i. eucl_down p (x $ i))" definition eucl_truncate_down_vec :: "nat \ ('a, 'b) vec \ ('a, 'b) vec" where "eucl_truncate_down_vec p x = (\ i. eucl_truncate_down p (x $ i))" definition eucl_truncate_up_vec :: "nat \ ('a, 'b) vec \ ('a, 'b) vec" where "eucl_truncate_up_vec p x = (\ i. eucl_truncate_up p (x $ i))" instance proof show *: "set (Basis_list::('a, 'b) vec list) = Basis" unfolding Basis_list_vec_def Basis_vec_def apply (auto simp: Basis_list_vec_def vec_eq_iff distinct_map Basis_vec_def intro!: distinct_concat inj_onI split: if_splits) apply (auto simp: Basis_list_vec_def vec_eq_iff distinct_map enum_distinct UNIV_enum[symmetric] intro!: distinct_concat inj_onI split: if_splits) done have "length (Basis_list::('a, 'b) vec list) = CARD('b) * DIM('a)" by (auto simp: Basis_list_vec_def length_concat o_def enum_distinct sum_list_distinct_conv_sum_set UNIV_enum[symmetric]) then show "distinct (Basis_list::('a, 'b) vec list)" using * by (auto intro!: card_distinct) qed (simp_all only: vector_cart[symmetric] vec_eq_iff eucl_down_vec_def eucl_down_def eucl_truncate_down_vec_def eucl_truncate_down_def eucl_truncate_up_vec_def eucl_truncate_up_def, auto simp: euclidean_vec_componentwise inner_axis Basis_list_vec_def vec_nth_inner_scaleR_craziness intro!: sum.cong[OF refl]) end lemma concat_same_lengths_nth: assumes "\xs. xs \ set XS \ length xs = N" assumes "i < length XS * N" "N > 0" shows "concat XS ! i = XS ! (i div N) ! (i mod N)" using assms apply (induction XS arbitrary: i) apply (auto simp: nth_append nth_Cons split: nat.splits) apply (simp add: div_eq_0_iff) by (metis Suc_inject div_geq mod_geq) lemma concat_map_map_index: shows "concat (map (\n. map (f n) xs) ys) = map (\i. f (ys ! (i div length xs)) (xs ! (i mod length xs))) [0..(x, y)\zip xs (map g xs). f x y) = (\x\set xs. f x (g x))" by (force simp add: sum_list_distinct_conv_sum_set assms distinct_zipI1 split_beta' in_set_zip in_set_conv_nth inj_on_convol_ident intro!: sum.reindex_cong[where l="\x. (x, g x)"]) lemma sum_list_zip_map_of: assumes "distinct bs" assumes "length xs = length bs" shows "(\(x, y)\zip xs bs. f x y) = (\x\set bs. f (the (map_of (zip bs xs) x)) x)" proof - have "(\(x, y)\zip xs bs. f x y) = (\(y, x)\zip bs xs. f x y)" by (subst zip_commute) (auto simp: o_def split_beta') also have "\ = (\(x, y)\zip bs (map (the o map_of (zip bs xs)) bs). f y x)" proof (rule arg_cong, rule map_cong) have "xs = (map (the \ map_of (zip bs xs)) bs)" using assms by (auto intro!: nth_equalityI simp: map_nth map_of_zip_nth) then show "zip bs xs = zip bs (map (the \ map_of (zip bs xs)) bs)" by simp qed auto also have "\ = (\x\set bs. f (the (map_of (zip bs xs) x)) x)" using assms(1) by (subst sum_list_zip_map) (auto simp: o_def) finally show ?thesis . qed lemma vec_nth_matrix: "vec_nth (vec_nth (matrix y) i) j = vec_nth (y (axis j 1)) i" unfolding matrix_def by simp lemma matrix_eqI: assumes "\x. x \ Basis \ A *v x = B *v x" shows "(A::real^'n^'n) = B" apply vector using assms apply (auto simp: Basis_vec_def) by (metis cart_eq_inner_axis matrix_vector_mul_component) lemma matrix_columnI: assumes "\i. column i A = column i B" shows "(A::real^'n^'n) = B" using assms apply vector apply (auto simp: column_def) apply vector by (metis iso_tuple_UNIV_I vec_lambda_inject) lemma vec_nth_Basis: fixes x::"real^'n" shows "x \ Basis \ vec_nth x i = (if x = axis i 1 then 1 else 0)" apply (auto simp: Basis_vec_def) by (metis cart_eq_inner_axis inner_axis_axis) lemma vec_nth_eucl_of_list_eq: "length M = CARD('n) \ vec_nth (eucl_of_list M::real^'n::enum) i = M ! index Basis_list (axis i (1::real))" apply (auto simp: eucl_of_list_def) apply (subst sum_list_zip_map_of) apply (auto intro!: distinct_zipI2 simp: split_beta') apply (subst sum.cong[OF refl]) apply (subst vec_nth_Basis) apply (force simp: set_zip) apply (rule refl) apply (auto simp: if_distrib sum.delta cong: if_cong) subgoal apply (cases "map_of (zip Basis_list M) (axis i 1::real^'n::enum)") subgoal premises prems proof - have "fst ` set (zip Basis_list M) = (Basis::(real^'n::enum) set)" using prems by (auto simp: in_set_zip) then show ?thesis using prems by (subst (asm) map_of_eq_None_iff) simp qed subgoal for a apply (auto simp: in_set_zip) subgoal premises prems for n by (metis DIM_cart DIM_real index_Basis_list_nth mult.right_neutral prems(2) prems(3)) done done done lemma index_Basis_list_axis1: "index Basis_list (axis i (1::real)) = index enum_class.enum i" apply (auto simp: Basis_list_vec_def Basis_list_real_def ) apply (subst index_map_inj) by (auto intro!: injI simp: axis_eq_axis) lemma vec_nth_eq_list_of_eucl1: "(vec_nth (M::real^'n::enum) i) = list_of_eucl M ! (index enum_class.enum i)" apply (subst eucl_of_list_list_of_eucl[of M, symmetric]) apply (subst vec_nth_eucl_of_list_eq) unfolding index_Basis_list_axis1 by auto lemma enum_3[simp]: "(enum_class.enum::3 list) = [0, 1, 2]" by code_simp+ lemma three_eq_zero: "(3::3) = 0" by simp lemma forall_3': "(\i::3. P i) \ P 0 \ P 1 \ P 2" using forall_3 three_eq_zero by auto lemma euclidean_eq_list_of_euclI: "x = y" if "list_of_eucl x = list_of_eucl y" using that by (metis eucl_of_list_list_of_eucl) lemma axis_one_neq_zero[simp]: "axis xa (1::'a::zero_neq_one) \ 0" by (auto simp: axis_def vec_eq_iff) lemma eucl_of_list_vec_nth3[simp]: "(eucl_of_list [g, h, i]::real^3) $ 0 = g" "(eucl_of_list [g, h, i]::real^3) $ 1 = h" "(eucl_of_list [g, h, i]::real^3) $ 2 = i" "(eucl_of_list [g, h, i]::real^3) $ 3 = g" by (auto simp: cart_eq_inner_axis eucl_of_list_inner vec_nth_eq_list_of_eucl1 index_Basis_list_axis1) type_synonym R3 = "real*real*real" lemma Basis_list_R3: "Basis_list = [(1,0,0), (0, 1, 0), (0, 0, 1)::R3]" by (auto simp: Basis_list_prod_def Basis_list_real_def zero_prod_def) lemma Basis_list_vec3: "Basis_list = [axis 0 1::real^3, axis 1 1, axis 2 1]" by (auto simp: Basis_list_vec_def Basis_list_real_def) lemma eucl_of_list3[simp]: "eucl_of_list [a, b, c] = (a, b, c)" by (auto simp: eucl_of_list_inner Basis_list_vec_def zero_prod_def Basis_prod_def Basis_list_vec3 Basis_list_R3 intro!: euclidean_eqI[where 'a=R3]) subsection \Bounded Linear Functions\ subsection \bounded linear functions\ locale blinfun_syntax begin no_notation vec_nth (infixl "$" 90) notation blinfun_apply (infixl "$" 999) end lemma bounded_linear_via_derivative: fixes f::"'a::real_normed_vector \ 'b::euclidean_space \\<^sub>L 'c::real_normed_vector" \ \TODO: generalize?\ assumes "\i. ((\x. blinfun_apply (f x) i) has_derivative (\x. f' y x i)) (at y)" shows "bounded_linear (f' y x)" proof - interpret linear "f' y x" proof (unfold_locales, goal_cases) case (1 v w) from has_derivative_unique[OF assms[of "v + w", unfolded blinfun.bilinear_simps] has_derivative_add[OF assms[of v] assms[of w]], THEN fun_cong, of x] show ?case . next case (2 r v) from has_derivative_unique[OF assms[of "r *\<^sub>R v", unfolded blinfun.bilinear_simps] has_derivative_scaleR_right[OF assms[of v], of r], THEN fun_cong, of x] show ?case . qed let ?bnd = "\i\Basis. norm (f' y x i)" { fix v have "f' y x v = (\i\Basis. (v \ i) *\<^sub>R f' y x i)" by (subst euclidean_representation[symmetric]) (simp add: sum scaleR) also have "norm \ \ norm v * ?bnd" by (auto intro!: order.trans[OF norm_sum] sum_mono mult_right_mono simp: sum_distrib_left Basis_le_norm) finally have "norm (f' y x v) \ norm v * ?bnd" . } then show ?thesis by unfold_locales auto qed definition blinfun_scaleR::"('a::real_normed_vector \\<^sub>L real) \ 'b::real_normed_vector \ ('a \\<^sub>L 'b)" where "blinfun_scaleR a b = blinfun_scaleR_left b o\<^sub>L a" lemma blinfun_scaleR_transfer[transfer_rule]: "rel_fun (pcr_blinfun (=) (=)) (rel_fun (=) (pcr_blinfun (=) (=))) (\a b c. a c *\<^sub>R b) blinfun_scaleR" by (auto simp: blinfun_scaleR_def rel_fun_def pcr_blinfun_def cr_blinfun_def OO_def) lemma blinfun_scaleR_rep_eq[simp]: "blinfun_scaleR a b c = a c *\<^sub>R b" by (simp add: blinfun_scaleR_def) lemma bounded_linear_blinfun_scaleR: "bounded_linear (blinfun_scaleR a)" unfolding blinfun_scaleR_def[abs_def] by (auto intro!: bounded_linear_intros) lemma blinfun_scaleR_has_derivative[derivative_intros]: assumes "(f has_derivative f') (at x within s)" shows "((\x. blinfun_scaleR a (f x)) has_derivative (\x. blinfun_scaleR a (f' x))) (at x within s)" using bounded_linear_blinfun_scaleR assms by (rule bounded_linear.has_derivative) lemma blinfun_componentwise: fixes f::"'a::real_normed_vector \ 'b::euclidean_space \\<^sub>L 'c::real_normed_vector" shows "f = (\x. \i\Basis. blinfun_scaleR (blinfun_inner_left i) (f x i))" by (auto intro!: blinfun_eqI simp: blinfun.sum_left euclidean_representation blinfun.scaleR_right[symmetric] blinfun.sum_right[symmetric]) lemma blinfun_has_derivative_componentwiseI: fixes f::"'a::real_normed_vector \ 'b::euclidean_space \\<^sub>L 'c::real_normed_vector" assumes "\i. i \ Basis \ ((\x. f x i) has_derivative blinfun_apply (f' i)) (at x)" shows "(f has_derivative (\x. \i\Basis. blinfun_scaleR (blinfun_inner_left i) (f' i x))) (at x)" by (subst blinfun_componentwise) (force intro: derivative_eq_intros assms simp: blinfun.bilinear_simps) lemma has_derivative_BlinfunI: fixes f::"'a::real_normed_vector \ 'b::euclidean_space \\<^sub>L 'c::real_normed_vector" assumes "\i. ((\x. f x i) has_derivative (\x. f' y x i)) (at y)" shows "(f has_derivative (\x. Blinfun (f' y x))) (at y)" proof - have 1: "f = (\x. \i\Basis. blinfun_scaleR (blinfun_inner_left i) (f x i))" by (rule blinfun_componentwise) moreover have 2: "(\ has_derivative (\x. \i\Basis. blinfun_scaleR (blinfun_inner_left i) (f' y x i))) (at y)" by (force intro: assms derivative_eq_intros) moreover interpret f': bounded_linear "f' y x" for x by (rule bounded_linear_via_derivative) (rule assms) have 3: "(\i\Basis. blinfun_scaleR (blinfun_inner_left i) (f' y x i)) i = f' y x i" for x i by (auto simp: if_distrib if_distribR blinfun.bilinear_simps f'.scaleR[symmetric] f'.sum[symmetric] euclidean_representation intro!: blinfun_euclidean_eqI) have 4: "blinfun_apply (Blinfun (f' y x)) = f' y x" for x apply (subst bounded_linear_Blinfun_apply) subgoal by unfold_locales subgoal by simp done show ?thesis apply (subst 1) apply (rule 2[THEN has_derivative_eq_rhs]) apply (rule ext) apply (rule blinfun_eqI) apply (subst 3) apply (subst 4) apply (rule refl) done qed lemma has_derivative_Blinfun: assumes "(f has_derivative f') F" shows "(f has_derivative Blinfun f') F" using assms by (subst bounded_linear_Blinfun_apply) auto lift_definition flip_blinfun:: "('a::real_normed_vector \\<^sub>L 'b::real_normed_vector \\<^sub>L 'c::real_normed_vector) \ 'b \\<^sub>L 'a \\<^sub>L 'c" is "\f x y. f y x" using bounded_bilinear.bounded_linear_left bounded_bilinear.bounded_linear_right bounded_bilinear.flip by auto lemma flip_blinfun_apply[simp]: "flip_blinfun f a b = f b a" by transfer simp lemma le_norm_blinfun: shows "norm (blinfun_apply f x) / norm x \ norm f" by transfer (rule le_onorm) lemma norm_flip_blinfun[simp]: "norm (flip_blinfun x) = norm x" (is "?l = ?r") proof (rule antisym) from order_trans[OF norm_blinfun, OF mult_right_mono, OF norm_blinfun, OF norm_ge_zero, of x] show "?l \ ?r" by (auto intro!: norm_blinfun_bound simp: ac_simps) have "norm (x a b) \ norm (flip_blinfun x) * norm a * norm b" for a b proof - have "norm (x a b) / norm a \ norm (flip_blinfun x b)" by (rule order_trans[OF _ le_norm_blinfun]) auto also have "\ \ norm (flip_blinfun x) * norm b" by (rule norm_blinfun) finally show ?thesis by (auto simp add: divide_simps blinfun.bilinear_simps algebra_simps split: if_split_asm) qed then show "?r \ ?l" by (auto intro!: norm_blinfun_bound) qed lemma bounded_linear_flip_blinfun[bounded_linear]: "bounded_linear flip_blinfun" by unfold_locales (auto simp: blinfun.bilinear_simps intro!: blinfun_eqI exI[where x=1]) lemma dist_swap2_swap2[simp]: "dist (flip_blinfun f) (flip_blinfun g) = dist f g" by (metis (no_types) bounded_linear_flip_blinfun dist_blinfun_def linear_simps(2) norm_flip_blinfun) context includes blinfun.lifting begin lift_definition blinfun_of_vmatrix::"(real^'m^'n) \ ((real^('m::finite)) \\<^sub>L (real^('n::finite)))" is "matrix_vector_mult:: ((real, 'm) vec, 'n) vec \ ((real, 'm) vec \ (real, 'n) vec)" unfolding linear_linear by (rule matrix_vector_mul_linear) lemma matrix_blinfun_of_vmatrix[simp]: "matrix (blinfun_of_vmatrix M) = M" apply vector apply (auto simp: matrix_def) apply transfer by (metis cart_eq_inner_axis matrix_vector_mul_component) end lemma blinfun_apply_componentwise: "B = (\i\Basis. blinfun_scaleR (blinfun_inner_left i) (blinfun_apply B i))" using blinfun_componentwise[of "\x. B", unfolded fun_eq_iff] by blast lemma blinfun_apply_eq_sum: assumes [simp]: "length v = CARD('n)" shows "blinfun_apply (B::(real^'n::enum)\\<^sub>L(real^'m::enum)) (eucl_of_list v) = (\ij Basis_list ! i) * v ! j) *\<^sub>R (Basis_list ! i))" apply (subst blinfun_apply_componentwise[of B]) apply (auto intro!: euclidean_eqI[where 'a="(real,'m) vec"] simp: blinfun.bilinear_simps eucl_of_list_inner inner_sum_left inner_Basis if_distrib sum_Basis_sum_nth_Basis_list nth_eq_iff_index if_distribR cong: if_cong) apply (subst sum.swap) by (auto simp: sum.delta algebra_simps) lemma in_square_lemma[intro, simp]: "x * C + y < D * C" if "x < D" "y < C" for x::nat proof - have "x * C + y < (D - 1) * C + C" apply (rule add_le_less_mono) apply (rule mult_right_mono) using that by auto also have "\ \ D * C" using that by (auto simp: algebra_simps) finally show ?thesis . qed lemma less_square_imp_div_less[intro, simp]: "i < E * D \ i div E < D" for i::nat by (metis div_eq_0_iff div_mult2_eq gr_implies_not0 mult_not_zero) lemma in_square_lemma'[intro, simp]: "i < L \ n < N \ i * N + n < N * L" for i n::nat by (metis in_square_lemma mult.commute) lemma distinct_nth_eq_iff: "distinct xs \ x < length xs \ y < length xs \ xs ! x = xs ! y \ x = y" by (drule inj_on_nth[where I="{..i. axis (enum_class.enum ! (i div CARD('i))) (axis (enum_class.enum ! (i mod CARD('i))) 1)) ((index enum_class.enum j) * CARD('i) + index enum_class.enum i)" by (auto simp: index_less_cardi enum_UNIV) note less=in_square_lemma[OF index_less_cardj index_less_cardi, of j i] show ?thesis apply (subst *) apply (subst index_map_inj_on[where S="{.. Basis \ vec_nth (vec_nth x i) j = ((if x = axis i (axis j 1) then 1 else 0))" by (auto simp: Basis_vec_def axis_def) lemma vec_nth_eucl_of_list_eq2: "length M = CARD('n) * CARD('m) \ vec_nth (vec_nth (eucl_of_list M::real^'n::enum^'m::enum) i) j = M ! index Basis_list (axis i (axis j (1::real)))" apply (auto simp: eucl_of_list_def) apply (subst sum_list_zip_map_of) apply (auto intro!: distinct_zipI2 simp: split_beta') apply (subst sum.cong[OF refl]) apply (subst vec_nth_Basis2) apply (force simp: set_zip) apply (rule refl) apply (auto simp: if_distrib sum.delta cong: if_cong) subgoal apply (cases "map_of (zip Basis_list M) (axis i (axis j 1)::real^'n::enum^'m::enum)") subgoal premises prems proof - have "fst ` set (zip Basis_list M) = (Basis::(real^'n::enum^'m::enum) set)" using prems by (auto simp: in_set_zip) then show ?thesis using prems by (subst (asm) map_of_eq_None_iff) auto qed subgoal for a apply (auto simp: in_set_zip) subgoal premises prems for n proof - have "n < card (Basis::(real^'n::_^'m::_) set)" by (simp add: prems(4)) then show ?thesis by (metis index_Basis_list_nth prems(2)) qed done done done lemma vec_nth_eq_list_of_eucl2: "vec_nth (vec_nth (M::real^'n::enum^'m::enum) i) j = list_of_eucl M ! (index enum_class.enum i * CARD('n) + index enum_class.enum j)" apply (subst eucl_of_list_list_of_eucl[of M, symmetric]) apply (subst vec_nth_eucl_of_list_eq2) unfolding index_Basis_list_axis2 by auto theorem eucl_of_list_matrix_vector_mult_eq_sum_nth_Basis_list: assumes "length M = CARD('n) * CARD('m)" assumes "length v = CARD('n)" shows "(eucl_of_list M::real^'n::enum^'m::enum) *v eucl_of_list v = (\ijR Basis_list ! i)" apply (vector matrix_vector_mult_def) apply (auto simp: ) apply (subst vec_nth_eucl_of_list_eq2) apply (auto simp: assms) apply (subst vec_nth_eucl_of_list_eq) apply (auto simp: assms index_Basis_list_axis2 index_Basis_list_axis1 vec_nth_Basis sum.delta nth_eq_iff_index if_distrib cong: if_cong) subgoal for i apply (rule sum.reindex_cong[where l="nth enum_class.enum"]) apply (auto simp: enum_distinct card_UNIV_length_enum distinct_nth_eq_iff intro!: inj_onI) apply (rule image_eqI[OF ]) apply (rule nth_index[symmetric]) apply (auto simp: enum_UNIV) by (auto simp: algebra_simps enum_UNIV enum_distinct index_nth_id) subgoal for i using index_less[of i "enum_class.enum" "CARD('n)"] by (auto simp: enum_UNIV card_UNIV_length_enum) done lemma index_enum_less[intro, simp]: "index enum_class.enum (i::'n::enum) < CARD('n)" by (auto intro!: index_less simp: enum_UNIV card_UNIV_length_enum) lemmas [intro, simp] = enum_distinct lemmas [simp] = card_UNIV_length_enum[symmetric] enum_UNIV lemma sum_index_enum_eq: "(\(k::'n::enum)\UNIV. f (index enum_class.enum k)) = (\iSubadditive and submultiplicative sequences\ theory Fekete - imports "HOL-Analysis.Ordered_Euclidean_Space" + imports "HOL-Analysis.Multivariate_Analysis" begin text \A real sequence is subadditive if $u_{n+m} \leq u_n+u_m$. This implies the convergence of $u_n/n$ to $Inf\{u_n/n\} \in [-\infty, +\infty)$, a useful result known as Fekete lemma. We prove it below. Taking logarithms, the same result applies to submultiplicative sequences. We illustrate it with the definition of the spectral radius as the limit of $\|x^n\|^{1/n}$, the convergence following from Fekete lemma.\ subsection \Subadditive sequences\ text \We define subadditive sequences, either from the start or eventually.\ definition subadditive::"(nat\real) \ bool" where "subadditive u = (\m n. u (m+n) \ u m + u n)" lemma subadditiveI: assumes "\m n. u (m+n) \ u m + u n" shows "subadditive u" unfolding subadditive_def using assms by auto lemma subadditiveD: assumes "subadditive u" shows "u (m+n) \ u m + u n" using assms unfolding subadditive_def by auto lemma subadditive_un_le_nu1: assumes "subadditive u" "n > 0" shows "u n \ n * u 1" proof - have *: "n = 0 \ (u n \ n * u 1)" for n proof (induction n) case 0 then show ?case by auto next case (Suc n) consider "n = 0" | "n > 0" by auto then show ?case proof (cases) case 1 then show ?thesis by auto next case 2 then have "u (Suc n) \ u n + u 1" using subadditiveD[OF assms(1), of n 1] by auto then show ?thesis using Suc.IH 2 by (auto simp add: algebra_simps) qed qed show ?thesis using *[of n] \n > 0\ by auto qed definition eventually_subadditive::"(nat\real) \ nat \ bool" where "eventually_subadditive u N0 = (\m>N0. \n>N0. u (m+n) \ u m + u n)" lemma eventually_subadditiveI: assumes "\m n. m > N0 \ n > N0 \ u (m+n) \ u m + u n" shows "eventually_subadditive u N0" unfolding eventually_subadditive_def using assms by auto lemma subadditive_imp_eventually_subadditive: assumes "subadditive u" shows "eventually_subadditive u 0" using assms unfolding subadditive_def eventually_subadditive_def by auto text \The main inequality that will lead to convergence is given in the next lemma: given $n$, then eventually $u_m/m$ is bounded by $u_n/n$, up to an arbitrarily small error. This is proved by doing the euclidean division of $m$ by $n$ and using the subadditivity. (the remainder in the euclidean division will give the error term.)\ lemma eventually_subadditive_ineq: assumes "eventually_subadditive u N0" "e>0" "n>N0" shows "\N>N0. \m\N. u m/m < u n/n + e" proof - have ineq_rec: "u(a*n+r) \ a * u n + u r" if "n>N0" "r>N0" for a n r proof (induct a) case (Suc a) have "a*n+r>N0" using \r>N0\ by simp have "u((Suc a)*n+r) = u(a*n+r+n)" by (simp add: algebra_simps) also have "... \ u(a*n+r)+u n" using assms \n>N0\ \a*n+r>N0\ eventually_subadditive_def by blast also have "... \ a*u n + u r + u n" by (simp add: Suc.hyps) also have "... = (Suc a) * u n + u r" by (simp add: algebra_simps) finally show ?case by simp qed (simp) have "n>0" "real n > 0" using \n>N0\ by auto define C where "C = Max {abs(u i) |i. i\2*n}" have ineq_C: "abs(u i) \ C" if "i \ 2 * n" for i unfolding C_def by (intro Max_ge, auto simp add: that) have ineq_all_m: "u m/m \ u n/n + 3*C/m" if "m\n" for m proof - have "real m>0" using \m\n\ \0 < real n\ by linarith obtain a0 r0 where "r00 < n\ mod_div_decomp mod_less_divisor by blast define a where "a = a0-1" define r where "r = r0+n" have "r<2*n" "r\n" unfolding r_def by (auto simp add: \r0) have "a0>0" using \m = a0*n + r0\ \n \ m\ \r0 < n\ not_le by fastforce then have "m = a * n + r" using a_def r_def \m = a0*n+r0\ mult_eq_if by auto then have real_eq: "-r = real n * a - m" by simp have "r>N0" using \r\n\ \n>N0\ by simp then have "u m \ a * u n + u r" using ineq_rec \m = a*n+r\ \n>N0\ by simp then have "n * u m \ n * (a * u n + u r)" using \real n>0\ by simp then have "n * u m - m * u n \ -r * u n + n * u r" unfolding real_eq by (simp add: algebra_simps) also have "... \ r * abs(u n) + n * abs(u r)" apply (intro add_mono mult_left_mono) using real_0_le_add_iff by fastforce+ also have "... \ (2 * n) * C + n * C" apply (intro add_mono mult_mono ineq_C) using less_imp_le[OF \r < 2 * n\] by auto finally have "n * u m - m * u n \ 3*C*n" by auto then show "u m/m \ u n/n + 3*C/m" using \0 < real n\ \0 < real m\ by (simp add: divide_simps mult.commute) qed obtain M::nat where M: "M \ 3 * C / e" using real_nat_ceiling_ge by auto define N where "N = M + n + N0 + 1" have "N > 3 * C / e" "N \ n" "N > N0" unfolding N_def using M by auto have "u m/m < u n/n + e" if "m \ N" for m proof - have "3 * C / m < e" using that \N > 3 * C / e\ \e > 0\ apply (auto simp add: algebra_simps divide_simps) by (meson le_less_trans linorder_not_le mult_less_cancel_left_pos of_nat_less_iff) then show ?thesis using ineq_all_m[of m] \n \ N\ \N \ m\ by auto qed then show ?thesis using \N0 < N\ by blast qed text \From the inequality above, we deduce the convergence of $u_n/n$ to its infimum. As this infimum might be $-\infty$, we formulate this convergence in the extended reals. Then, we specialize it to the real situation, separating the cases where $u_n/n$ is bounded below or not.\ lemma subadditive_converges_ereal': assumes "eventually_subadditive u N0" shows "(\m. ereal(u m/m)) \ Inf {ereal(u n/n) | n. n>N0}" proof - define v where "v = (\m. ereal(u m/m))" define V where "V = {v n | n. n>N0}" define l where "l = Inf V" have "\t. t\V \ t\l" by (simp add: Inf_lower l_def) then have "v n \ l" if "n > N0" for n using V_def that by blast then have lower: "eventually (\n. a < v n) sequentially" if "a < l" for a by (meson that dual_order.strict_trans1 eventually_at_top_dense) have upper: "eventually (\n. a > v n) sequentially" if "a > l" for a proof - obtain t where "t\V" "ta>l\ Inf_greatest l_def not_le) then obtain e::real where "e>0" "t+e < a" by (meson ereal_le_epsilon2 leD le_less_linear) obtain n where "n>N0" "t = u n/n" using V_def v_def \t \ V\ by blast then have "u n/n + e < a" using \t+e < a\ by simp obtain N where "\m\N. u m/m < u n/n + e" using eventually_subadditive_ineq[OF assms] \0 < e\ \N0 < n\ by blast then have "u m/m < a" if "m \ N" for m using that \u n/n + e < a\ less_ereal.simps(1) less_trans by blast then have "v m< a" if "m \ N" for m using v_def that by blast then show ?thesis using eventually_at_top_linorder by auto qed show ?thesis using lower upper unfolding V_def l_def v_def by (simp add: order_tendsto_iff) qed lemma subadditive_converges_ereal: assumes "subadditive u" shows "(\m. ereal(u m/m)) \ Inf {ereal(u n/n) | n. n>0}" by (rule subadditive_converges_ereal'[OF subadditive_imp_eventually_subadditive[OF assms]]) lemma subadditive_converges_bounded': assumes "eventually_subadditive u N0" "bdd_below {u n/n | n. n>N0}" shows "(\n. u n/n) \ Inf {u n/n | n. n>N0}" proof- have *: "(\n. ereal(u n /n)) \ Inf {ereal(u n/n)|n. n > N0}" by (simp add: assms(1) subadditive_converges_ereal') define V where "V = {u n/n | n. n>N0}" have a: "bdd_below V" "V\{}" by (auto simp add: V_def assms(2)) have "Inf {ereal(t)| t. t\V} = ereal(Inf V)" by (subst ereal_Inf'[OF a], simp add: Setcompr_eq_image) moreover have "{ereal(t)| t. t\V} = {ereal(u n/n)|n. n > N0}" using V_def by blast ultimately have "Inf {ereal(u n/n)|n. n > N0} = ereal(Inf {u n/n |n. n > N0})" using V_def by auto then have "(\n. ereal(u n /n)) \ ereal(Inf {u n/n | n. n>N0})" using * by auto then show ?thesis by simp qed lemma subadditive_converges_bounded: assumes "subadditive u" "bdd_below {u n/n | n. n>0}" shows "(\n. u n/n) \ Inf {u n/n | n. n>0}" by (rule subadditive_converges_bounded'[OF subadditive_imp_eventually_subadditive[OF assms(1)] assms(2)]) text \We reformulate the previous lemma in a more directly usable form, avoiding the infimum.\ lemma subadditive_converges_bounded'': assumes "subadditive u" "\n. n > 0 \ u n \ n * (a::real)" shows "\l. (\n. u n / n) \ l \ (\n>0. u n \ n * l)" proof - have B: "bdd_below {u n/n | n. n>0}" apply (rule bdd_belowI[of _ a]) using assms(2) apply (auto simp add: divide_simps) apply (metis mult.commute mult_left_le_imp_le of_nat_0_less_iff) done define l where "l = Inf {u n/n | n. n>0}" have *: "u n / n \ l" if "n > 0" for n unfolding l_def using that by (auto intro!: cInf_lower[OF _ B]) show ?thesis apply (rule exI[of _ l], auto) using subadditive_converges_bounded[OF assms(1) B] apply (simp add: l_def) using * by (simp add: divide_simps algebra_simps) qed lemma subadditive_converges_unbounded': assumes "eventually_subadditive u N0" "\ (bdd_below {u n/n | n. n>N0})" shows "(\n. ereal(u n/n)) \ -\" proof - have *: "(\n. ereal(u n /n)) \ Inf {ereal(u n/n)|n. n > N0}" by (simp add: assms(1) subadditive_converges_ereal') define V where "V = {u n/n | n. n>N0}" then have "\ bdd_below V" using assms by simp have "Inf {ereal(t) | t. t\V} = -\" by (rule ereal_bot, metis (mono_tags, lifting) \\ bdd_below V\ bdd_below_def leI Inf_lower2 ereal_less_eq(3) le_less mem_Collect_eq) moreover have "{ereal(t)| t. t\V} = {ereal(u n/n)|n. n > N0}" using V_def by blast ultimately have "Inf {ereal(u n/n)|n. n > N0} = -\" by auto then show ?thesis using * by simp qed lemma subadditive_converges_unbounded: assumes "subadditive u" "\ (bdd_below {u n/n | n. n>0})" shows "(\n. ereal(u n/n)) \ -\" by (rule subadditive_converges_unbounded'[OF subadditive_imp_eventually_subadditive[OF assms(1)] assms(2)]) subsection \Superadditive sequences\ text \While most applications involve subadditive sequences, one sometimes encounters superadditive sequences. We reformulate quickly some of the above results in this setting.\ definition superadditive::"(nat\real) \ bool" where "superadditive u = (\m n. u (m+n) \ u m + u n)" lemma subadditive_of_superadditive: assumes "superadditive u" shows "subadditive (\n. -u n)" using assms unfolding superadditive_def subadditive_def by (auto simp add: algebra_simps) lemma superadditive_un_ge_nu1: assumes "superadditive u" "n > 0" shows "u n \ n * u 1" using subadditive_un_le_nu1[OF subadditive_of_superadditive[OF assms(1)] assms(2)] by auto lemma superadditive_converges_bounded'': assumes "superadditive u" "\n. n > 0 \ u n \ n * (a::real)" shows "\l. (\n. u n / n) \ l \ (\n>0. u n \ n * l)" proof - have "\l. (\n. -u n / n) \ l \ (\n>0. -u n \ n * l)" apply (rule subadditive_converges_bounded''[OF subadditive_of_superadditive[OF assms(1)], of "-a"]) using assms(2) by auto then obtain l where l: "(\n. -u n / n) \ l" "(\n>0. -u n \ n * l)" by blast have "(\n. -((-u n)/n)) \ -l" by (intro tendsto_intros l) moreover have "\n>0. u n \ n * (-l)" using l(2) by (auto simp add: algebra_simps) (metis minus_equation_iff neg_le_iff_le) ultimately show ?thesis by auto qed subsection \Almost additive sequences\ text \One often encounters sequences which are both subadditive and superadditive, but only up to an additive constant. Adding or subtracting this constant, one can make the sequence genuinely subadditive or superadditive, and thus deduce results about its convergence, as follows. Such sequences appear notably when dealing with quasimorphisms.\ lemma almost_additive_converges: fixes u::"nat \ real" assumes "\m n. abs(u(m+n) - u m - u n) \ C" shows "convergent (\n. u n/n)" "abs(u k - k * lim (\n. u n / n)) \ C" proof - have "(abs (u 0)) \ C" using assms[of 0 0] by auto then have "C \ 0" by auto define v where "v = (\n. u n + C)" have "subadditive v" unfolding subadditive_def v_def using assms by (auto simp add: algebra_simps abs_diff_le_iff) then have vle: "v n \ n * v 1" if "n > 0" for n using subadditive_un_le_nu1 that by auto define w where "w = (\n. u n - C)" have "superadditive w" unfolding superadditive_def w_def using assms by (auto simp add: algebra_simps abs_diff_le_iff) then have wge: "w n \ n * w 1" if "n > 0" for n using superadditive_un_ge_nu1 that by auto have I: "v n \ w n" for n unfolding v_def w_def using \C \ 0\ by auto then have *: "v n \ n * w 1" if "n > 0" for n using order_trans[OF wge[OF that]] by auto then obtain lv where lv: "(\n. v n/n) \ lv" "\n. n > 0 \ v n \ n * lv" using subadditive_converges_bounded''[OF \subadditive v\ *] by auto have *: "w n \ n * v 1" if "n > 0" for n using order_trans[OF _ vle[OF that]] I by auto then obtain lw where lw: "(\n. w n/n) \ lw" "\n. n > 0 \ w n \ n * lw" using superadditive_converges_bounded''[OF \superadditive w\ *] by auto have *: "v n/n = w n /n + 2*C*(1/n)" for n unfolding v_def w_def by (auto simp add: algebra_simps divide_simps) have "(\n. w n /n + 2*C*(1/n)) \ lw + 2*C*0" by (intro tendsto_add tendsto_mult lim_1_over_n lw, auto) then have "lw = lv" unfolding *[symmetric] using lv(1) LIMSEQ_unique by auto have *: "u n/n = w n /n + C*(1/n)" for n unfolding w_def by (auto simp add: algebra_simps divide_simps) have "(\n. u n /n) \ lw + C*0" unfolding * by (intro tendsto_add tendsto_mult lim_1_over_n lw, auto) then have lu: "convergent (\n. u n/n)" "lim (\n. u n/n) = lw" by (auto simp add: convergentI limI) then show "convergent (\n. u n/n)" by simp show "abs(u k - k * lim (\n. u n / n)) \ C" proof (cases "k>0") case False then show ?thesis using assms[of 0 0] by auto next case True have "u k - k * lim (\n. u n/n) = v k - C - k * lv" unfolding lu(2) \lw = lv\ v_def by auto also have "... \ -C" using lv(2)[OF True] by auto finally have A: "u k - k * lim (\n. u n/n) \ - C" by simp have "u k - k * lim (\n. u n/n) = w k + C - k * lw" unfolding lu(2) w_def by auto also have "... \ C" using lw(2)[OF True] by auto finally show ?thesis using A by auto qed qed subsection \Submultiplicative sequences, application to the spectral radius\ text \In the same way as subadditive sequences, one may define submultiplicative sequences. Essentially, a sequence is submultiplicative if its logarithm is subadditive. A difference is that we allow a submultiplicative sequence to take the value $0$, as this shows up in applications. This implies that we have to distinguish in the proofs the situations where the value $0$ is taken or not. In the latter situation, we can use directly the results from the subadditive case to deduce convergence. In the former situation, convergence to $0$ is obvious as the sequence vanishes eventually.\ lemma submultiplicative_converges: fixes u::"nat\real" assumes "\n. u n \ 0" "\m n. u (m+n) \ u m * u n" shows "(\n. root n (u n))\ Inf {root n (u n) | n. n>0}" proof - define v where "v = (\ n. root n (u n))" define V where "V = {v n | n. n>0}" then have "V \ {}" by blast have "t \ 0" if "t \ V" for t using that V_def v_def assms(1) by auto then have "Inf V \ 0" by (simp add: \V \ {}\ cInf_greatest) have "bdd_below V" by (meson \\t. t \ V \ 0 \ t\ bdd_below_def) show ?thesis proof cases assume "\n. u n = 0" then obtain n where "u n = 0" by auto then have "u m = 0" if "m \ n" for m by (metis that antisym_conv assms(1) assms(2) le_Suc_ex mult_zero_left) then have *: "v m = 0" if "m \ n" for m using v_def that by simp then have "v \ 0" using lim_explicit by force have "v (Suc n) \ V" using V_def by blast moreover have "v (Suc n) = 0" using * by auto ultimately have "Inf V \ 0" by (simp add: \bdd_below V\ cInf_lower) then have "Inf V = 0" using \0 \ Inf V\ by auto then show ?thesis using V_def v_def \v \ 0\ by auto next assume "\ (\n. u n = 0)" then have "u n > 0" for n by (metis assms(1) less_eq_real_def) define w where "w n = ln (u n)" for n have express_vn: "v n = exp(w n/n)" if "n>0" for n proof - have "(exp(w n/n))^n = exp(n*(w n/n))" by (metis exp_of_nat_mult) also have "... = exp(w n)" by (simp add: \0 < n\) also have "... = u n" by (simp add: \\n. 0 < u n\ w_def) finally have "exp(w n/n) = root n (u n)" by (metis \0 < n\ exp_ge_zero real_root_power_cancel) then show ?thesis unfolding v_def by simp qed have "eventually_subadditive w 0" proof (rule eventually_subadditiveI) fix m n have "w (m+n) = ln (u (m+n))" by (simp add: w_def) also have "... \ ln(u m * u n)" by (meson \\n. 0 < u n\ assms(2) zero_less_mult_iff ln_le_cancel_iff) also have "... = ln(u m) + ln(u n)" by (simp add: \\n. 0 < u n\ ln_mult) also have "... = w m + w n" by (simp add: w_def) finally show "w (m+n) \ w m + w n". qed define l where "l = Inf V" then have "v n\l" if "n > 0" for n using V_def that by (metis (mono_tags, lifting) \bdd_below V\ cInf_lower mem_Collect_eq) then have lower: "eventually (\n. a < v n) sequentially" if "a < l" for a by (meson that dual_order.strict_trans1 eventually_at_top_dense) have upper: "eventually (\n. a > v n) sequentially" if "a > l" for a proof - obtain t where "t\V" "t < a" using \V \ {}\ cInf_lessD l_def \a>l\ by blast then have "t > 0" using V_def \\n. 0 < u n\ v_def by auto then have "a/t > 1" using \t by simp define e where "e = ln(a/t)/2" have "e > 0" "e < ln(a/t)" unfolding e_def by (simp_all add: \1 < a / t\ ln_gt_zero) then have "exp(e) < a/t" by (metis \1 < a / t\ exp_less_cancel_iff exp_ln less_trans zero_less_one) obtain n where "n>0" "t = v n" using V_def v_def \t \ V\ by blast with \0 < t\ have "v n * exp(e) < a" using \exp(e) < a/t\ by (auto simp add: field_simps) obtain N where *: "N>0" "\m. m\N \ w m/m < w n/n + e" using eventually_subadditive_ineq[OF \eventually_subadditive w 0\] \0 < n\ \e>0\ by blast have "v m < a" if "m \ N" for m proof - have "m>0" using that \N>0\ by simp have "w m/m < w n/n + e" by (simp add: \N \ m\ *) then have "exp(w m/m) < exp(w n/n + e)" by simp also have "... = exp(w n/n) * exp(e)" by (simp add: mult_exp_exp) finally have "v m < v n * exp(e)" using express_vn \m>0\ \n>0\ by simp then show "v m < a" using \v n * exp(e) < a\ by simp qed then show ?thesis using eventually_at_top_linorder by auto qed show ?thesis using lower upper unfolding v_def l_def V_def by (simp add: order_tendsto_iff) qed qed text \An important application of submultiplicativity is to prove the existence of the spectral radius of a matrix, as the limit of $\|A^n\|^{1/n}$.\ definition spectral_radius::"'a::real_normed_algebra_1 \ real" where "spectral_radius x = Inf {root n (norm(x^n))| n. n>0}" lemma spectral_radius_aux: fixes x::"'a::real_normed_algebra_1" defines "V \ {root n (norm(x^n))| n. n>0}" shows "\t. t\V \ t \ spectral_radius x" "\t. t\V \ t \ 0" "bdd_below V" "V \ {}" "Inf V \ 0" proof - show "V\{}" using V_def by blast show *: "t \ 0" if "t \ V" for t using that unfolding V_def using real_root_pos_pos_le by auto then show "bdd_below V" by (meson bdd_below_def) then show "Inf V \ 0" by (simp add: \V \ {}\ * cInf_greatest) show "\t. t\V \ t \ spectral_radius x" by (metis (mono_tags, lifting) \bdd_below V\ assms cInf_lower spectral_radius_def) qed lemma spectral_radius_nonneg [simp]: "spectral_radius x \ 0" by (simp add: spectral_radius_aux(5) spectral_radius_def) lemma spectral_radius_upper_bound [simp]: "(spectral_radius x)^n \ norm(x^n)" proof (cases) assume "\(n = 0)" have "root n (norm(x^n)) \ spectral_radius x" using spectral_radius_aux \n \ 0\ by auto then show ?thesis by (metis \n \ 0\ spectral_radius_nonneg norm_ge_zero not_gr0 power_mono real_root_pow_pos2) qed (simp) lemma spectral_radius_limit: "(\n. root n (norm(x^n))) \ spectral_radius x" proof - have "norm(x^(m+n)) \ norm(x^m) * norm(x^n)" for m n by (simp add: power_add norm_mult_ineq) then show ?thesis unfolding spectral_radius_def using submultiplicative_converges by auto qed end (*of Fekete.thy*) diff --git a/thys/First_Welfare_Theorem/Preferences.thy b/thys/First_Welfare_Theorem/Preferences.thy --- a/thys/First_Welfare_Theorem/Preferences.thy +++ b/thys/First_Welfare_Theorem/Preferences.thy @@ -1,651 +1,651 @@ (* License: LGPL *) (* Author: Julian Parsert Author: Cezary Kaliszyk *) section \Preference Relations\ text \ Preferences modeled as a set of pairs \ theory Preferences imports - "HOL-Analysis.Ordered_Euclidean_Space" + "HOL-Analysis.Multivariate_Analysis" Syntax begin subsection \Basic Preference Relation\ text \ Basic preference relation locale with carrier and relation modeled as a set of pairs. \ locale preference = fixes carrier :: "'a set" fixes relation :: "'a relation" assumes not_outside: "(x,y) \ relation \ x \ carrier" and "(x,y) \ relation \ y \ carrier" assumes trans_refl: "preorder_on carrier relation" context preference begin abbreviation geq ("_ \ _" [51,51] 60) where "x \ y \ x \[relation] y" abbreviation str_gr ("_ \ _" [51,51] 60) where "x \ y \ x \ y \ \y \ x" abbreviation indiff ("_ \ _" [51,51] 60) where "x \ y \ x \ y \ y \ x" lemma reflexivity: "refl_on carrier relation" using preorder_on_def trans_refl by blast lemma transitivity: "trans relation" using preorder_on_def trans_refl by auto lemma indiff_trans [simp]: "x \ y \ y \ z \ x \ z" by (meson transE transitivity) end subsubsection \ Contour sets \ definition at_least_as_good :: "'a \ 'a set \ 'a relation \ 'a set" where "at_least_as_good x B P = {y \ B. y \[P] x }" definition no_better_than :: "'a \ 'a set \ 'a relation \ 'a set" where "no_better_than x B P = {y \ B. x \[P] y}" definition as_good_as :: "'a \ 'a set \ 'a relation \ 'a set" where "as_good_as x B P = {y \ B. x \[P] y}" lemma at_lst_asgd_ge: assumes "x \ at_least_as_good y B Pr" shows "x \[Pr] y" using assms at_least_as_good_def by fastforce lemma strict_contour_is_diff: "{a \ B. a \[Pr] y} = at_least_as_good y B Pr - as_good_as y B Pr" by(auto simp add: at_least_as_good_def as_good_as_def) lemma strict_countour_def [simp]: "(at_least_as_good y B Pr) - as_good_as y B Pr = {x \ B. x \[Pr] y}" by (simp add: as_good_as_def at_least_as_good_def strict_contour_is_diff) lemma at_least_as_goodD [dest]: assumes "z \ at_least_as_good y B Pr" shows "z \[Pr] y" using assms at_least_as_good_def by fastforce subsection \Rational Preference Relation\ text \ Rational preferences add totality to the basic preferences. \ locale rational_preference = preference + assumes total: "total_on carrier relation" begin lemma compl: "\x \ carrier . \y\ carrier . x \ y \ y \ x" by (metis refl_onD reflexivity total total_on_def) lemma strict_not_refl_weak [iff]: "x \ carrier \ y \ carrier \ \ (y \ x) \ x \ y" by (metis refl_onD reflexivity total total_on_def) lemma strict_trans [simp]: "x \ y \ y \ z \ x \ z" by (meson transE transitivity) lemma completeD [dest]: "x \ carrier \ y \ carrier \ x \ y \ x \ y \ y \ x" by blast lemma pref_in_at_least_as: assumes "x \ y" shows "x \ at_least_as_good y carrier relation" by (metis (no_types, lifting) CollectI assms at_least_as_good_def preference.not_outside preference_axioms) lemma worse_in_no_better: assumes "x \ y" shows "y \ no_better_than y carrier relation" by (metis (no_types, lifting) CollectI assms no_better_than_def preference_axioms preference_def strict_not_refl_weak) lemma strict_is_neg_transitive : assumes "x \ carrier \ y \ carrier \ z \ carrier" shows "x \ y \ x \ z \ z \ y" by (meson assms compl transE transitivity) lemma weak_is_transitive: assumes "x \ carrier \ y \ carrier \ z \ carrier" shows "x \ y \ y \ z \ x \ z" by (meson transD transitivity) lemma no_better_than_nonepty: assumes "carrier \ {}" shows "\x. x \ carrier \ (no_better_than x carrier relation) \ {}" by (metis (no_types, lifting) empty_iff mem_Collect_eq no_better_than_def refl_onD reflexivity) lemma no_better_subset_pref : assumes "x \ y" shows "no_better_than y carrier relation \ no_better_than x carrier relation" proof fix a assume "a \ no_better_than y carrier relation" then show "a \ no_better_than x carrier relation" by (metis (no_types, lifting) assms mem_Collect_eq no_better_than_def transE transitivity) qed lemma no_better_thansubset_rel : assumes "x \ carrier" and "y \ carrier" assumes "no_better_than y carrier relation \ no_better_than x carrier relation" shows "x \ y" proof - have "{a \ carrier. y \ a} \ {a \ carrier. x \ a}" by (metis (no_types) assms(3) no_better_than_def) then show ?thesis by (metis (mono_tags, lifting) Collect_mono_iff assms(2) compl) qed lemma nbt_nest : shows "(no_better_than y carrier relation \ no_better_than x carrier relation) \ (no_better_than x carrier relation \ no_better_than y carrier relation)" by (metis (no_types, lifting) CollectD compl no_better_subset_pref no_better_than_def not_outside subsetI) lemma at_lst_asgd_not_ge: assumes "carrier \ {}" assumes "x \ carrier" and "y \ carrier" assumes "x \ at_least_as_good y carrier relation" shows "\ x \ y" by (metis (no_types, lifting) CollectI assms(2) assms(4) at_least_as_good_def) lemma as_good_as_sameIff[iff]: "z \ as_good_as y carrier relation \ z \ y \ y \ z" by (metis (no_types, lifting) as_good_as_def mem_Collect_eq not_outside) lemma same_at_least_as_equal: assumes "z \ y" shows "at_least_as_good z carrier relation = at_least_as_good y carrier relation" (is "?az = ?ay") proof have "z \ carrier \ y \ carrier" by (meson assms refl_onD2 reflexivity) moreover have "\x \ carrier. x \ z \ x \ y" by (meson assms transD transitivity) ultimately show "?az \ ?ay" by (metis at_lst_asgd_ge at_lst_asgd_not_ge equals0D not_outside subsetI) next have "z \ carrier \ y \ carrier" by (meson assms refl_onD2 reflexivity) moreover have "\x \ carrier. x \ y \ x \ z" by (meson assms transD transitivity) ultimately show "?ay \ ?az" by (metis at_lst_asgd_ge at_lst_asgd_not_ge equals0D not_outside subsetI) qed lemma as_good_asIff [iff]: "x \ as_good_as y carrier relation \ x \[relation] y" by blast lemma nbt_subset: assumes "finite carrier" assumes "x \ carrier" and "y \ carrier" shows "no_better_than x carrier relation \ no_better_than x carrier relation \ no_better_than x carrier relation \ no_better_than x carrier relation" by auto lemma fnt_carrier_fnt_rel: "finite carrier \ finite relation" by (metis finite_SigmaI refl_on_def reflexivity rev_finite_subset) lemma nbt_subset_carrier: assumes "x \ carrier" shows "no_better_than x carrier relation \ carrier" using no_better_than_def by fastforce lemma xy_in_eachothers_nbt: assumes "x \ carrier" "y \ carrier" shows "x \ no_better_than y carrier relation \ y \ no_better_than x carrier relation" by (meson assms(1) assms(2) contra_subsetD nbt_nest refl_onD reflexivity worse_in_no_better) lemma same_nbt_same_pref: assumes "x \ carrier" "y \ carrier" shows "x \ no_better_than y carrier relation \ y \ no_better_than x carrier relation \ x \ y" by (metis (mono_tags, lifting) CollectD contra_subsetD no_better_subset_pref no_better_than_def worse_in_no_better) lemma indifferent_imp_weak_pref: assumes "x \ y" shows" x \ y" "y \ x" by (simp add: assms)+ subsection \ Finite carrier\ context assumes "finite carrier" begin lemma fnt_carrier_fnt_nbt: shows "\x\carrier. finite (no_better_than x carrier relation)" proof fix x assume "x \ carrier" then show "finite (no_better_than x carrier relation)" using finite_subset nbt_subset_carrier \finite carrier\ by blast qed lemma nbt_subset_imp_card_leq: assumes "x \ carrier" and "y \ carrier" shows "no_better_than x carrier relation \ no_better_than y carrier relation \ card (no_better_than x carrier relation) \ card (no_better_than y carrier relation)" (is "?nbt \ ?card") proof assume "?nbt" then show "?card" by (simp add: assms(2) card_mono fnt_carrier_fnt_nbt) next assume "?card" then show "?nbt" by (metis assms(1) card_seteq fnt_carrier_fnt_nbt nbt_nest) qed lemma card_leq_pref: assumes "x \ carrier" and "y \ carrier" shows "card (no_better_than x carrier relation) \ card (no_better_than y carrier relation) \ y \ x" proof (rule iffI, goal_cases) case 1 then show ?case using assms(1) assms(2) nbt_subset_imp_card_leq no_better_thansubset_rel by presburger next case 2 then show ?case using assms(1) assms(2) nbt_subset_imp_card_leq no_better_subset_pref by blast qed lemma finite_ne_remove_induct: assumes "finite B" "B \ {}" "\A. finite A \ A \ B \ A \ {} \ (\x. x \ A \ A - {x} \ {} \ P (A - {x})) \ P A" shows "P B" by (metis assms finite_remove_induct[where P = "\F. F = {} \ P F" for P]) lemma finite_nempty_preorder_has_max: assumes "finite B" "B \ {}" "refl_on B R" "trans R" "total_on B R" shows "\x \ B. \y \ B. (x, y) \ R" using assms(1) subset_refl[of B] assms(2) proof (induct B rule: finite_subset_induct) case (insert x F) then show ?case using assms(3-) by (cases "F = {}") (auto simp: refl_onD total_on_def, metis refl_onD2 transE) qed auto lemma finite_nempty_preorder_has_min: assumes "finite B" "B \ {}" "refl_on B R" "trans R" "total_on B R" shows "\x \ B. \y \ B. (y, x) \ R" using assms(1) subset_refl[of B] assms(2) proof (induct B rule: finite_subset_induct) case (insert x F) then show ?case using assms(3-) by (cases "F = {}") (auto simp: refl_onD total_on_def, metis refl_onD2 transE) qed auto lemma finite_nonempty_carrier_has_maximum: assumes "carrier \ {}" shows "\e \ carrier. \m \ carrier. e \[relation] m" using finite_nempty_preorder_has_max[of carrier relation] assms \finite carrier\ reflexivity total transitivity by blast lemma finite_nonempty_carrier_has_minimum: assumes "carrier \ {}" shows "\e \ carrier. \m \ carrier. m \[relation] e" using finite_nempty_preorder_has_min[of carrier relation] assms \finite carrier\ reflexivity total transitivity by blast end (*finite carrier*) lemma all_carrier_ex_sub_rel: "\c \ carrier. \r \ relation. rational_preference c r" proof (standard,standard) fix c assume c_in: "c \ carrier" define r' where "r' = {(x,y) \ relation. x \ c \ y \ c}" have r'_sub: "r' \ c \ c" using r'_def by blast have "\x \ c. x \[r'] x" by (metis (no_types, lifting) CollectI c_in case_prodI compl r'_def subsetCE) then have refl: "refl_on c r'" by (meson r'_sub refl_onI) have trans: "trans r'" proof fix x y z assume a1: "x \[r'] y" assume a2: "y \[r'] z" show " x \[r'] z" by (metis (mono_tags, lifting) CollectD CollectI a1 a2 case_prodD case_prodI r'_def transE transitivity) qed have total: "total_on c r'" proof (standard) fix x y assume a1:"x \ c" assume a2: "y \ c" assume a3: "x \ y" show "x \[r'] y \ y \[r'] x " by (metis (no_types, lifting) CollectI a1 a2 c_in case_prodI compl r'_def subset_iff) qed have "rational_preference c r'" by (meson local.refl local.trans preference.intro preorder_on_def rational_preference.intro rational_preference_axioms.intro refl_on_domain total) thus "\r\relation. rational_preference c r" by (metis (no_types, lifting) CollectD case_prodD r'_def subrelI) qed end (*rational preference*) subsection \ Local Non-Satiation \ text \ Defining local non-satiation. \ definition local_nonsatiation where "local_nonsatiation B P \ (\x\B. \e>0. \y\B. norm (y - x) \ e \ y \[P] x)" text \ Alternate definitions and intro/dest rules with them \ lemma lns_alt_def1 [iff] : shows "local_nonsatiation B P \ (\x \ B. \e>0. (\y \ B. dist y x \ e \ y \[P] x))" by(simp add : dist_norm local_nonsatiation_def) lemma lns_normI [intro]: assumes "\x e. x \ B \ e > 0 \ (\y\B. norm (y - x) \ e \ y \[P] x)" shows "local_nonsatiation B P" by (simp add: assms dist_norm) lemma lns_distI [intro]: assumes "\x e. x \ B \ e > 0 \ (\y\B. (dist y x) \ e \ y \[P] x)" shows "local_nonsatiation B P" by (meson lns_alt_def1 assms) lemma lns_alt_def2 [iff]: "local_nonsatiation B P \ (\x\B. \e>0. (\y. y \ (ball x e) \ y \ B \ y \[P] x))" proof assume "local_nonsatiation B P" then show "\x\B. \e>0. \x'. x' \ ball x e \ x' \ B \ x' \[P] x" by (auto simp add : ball_def) (metis dense le_less_trans dist_commute) next assume "\x\B. \e>0. \y. y \ ball x e \ y \ B \ y \[P] x" then show "local_nonsatiation B P" by (metis (no_types, lifting) ball_def dist_commute less_le_not_le lns_alt_def1 mem_Collect_eq) qed lemma lns_normD [dest]: assumes "local_nonsatiation B P" shows "\x \ B. \e>0. \y \ B. (norm (y - x) \ e \ y \[P] x)" by (meson assms local_nonsatiation_def) subsection \ Convex preferences \ definition weak_convex_pref :: "('a::real_vector) relation \ bool" where "weak_convex_pref Pr \ (\x y. x \[Pr] y \ (\\ \. \ + \ = 1 \ \ > 0 \ \ > 0 \ \ *\<^sub>R x + \ *\<^sub>R y \[Pr] y))" definition convex_pref :: "('a::real_vector) relation \ bool" where "convex_pref Pr \ (\x y. x \[Pr] y \ (\\. 1 > \ \ \ > 0 \ \ *\<^sub>R x + (1-\) *\<^sub>R y \[Pr] y))" definition strict_convex_pref :: "('a::real_vector) relation \ bool" where "strict_convex_pref Pr \ (\x y. x \[Pr] y \ x \ y \ (\\. 1 > \ \ \ > 0 \ \ *\<^sub>R x + (1-\) *\<^sub>R y \[Pr] y))" lemma convex_ge_imp_conved: assumes "\x y. x \[Pr] y \ (\\ \. \ + \ = 1 \ \ \ 0 \ \ \ 0 \ \ *\<^sub>R x + \ *\<^sub>R y \[Pr] y)" shows "weak_convex_pref Pr" by (simp add: assms weak_convex_pref_def) lemma weak_convexI [intro]: assumes "\x y \ \. x \[Pr] y \ \ + \ = 1 \ 0 < \ \ 0 < \ \ \ *\<^sub>R x + \ *\<^sub>R y \[Pr] y" shows "weak_convex_pref Pr" by (simp add: assms weak_convex_pref_def) lemma weak_convexD [dest]: assumes "weak_convex_pref Pr" and "x \[Pr] y" and "0 < u" and "0 < v" and "u + v = 1" shows "u *\<^sub>R x + v *\<^sub>R y \[Pr] y" using assms weak_convex_pref_def by blast subsection \ Real Vector Preferences \ text \ Preference relations on real vector type class. \ locale real_vector_rpr = rational_preference carrier relation for carrier :: "'a::real_vector set" and relation :: "'a relation" sublocale real_vector_rpr \ rational_preference carrier relation by (simp add: rational_preference_axioms) context real_vector_rpr begin lemma have_rpr: "rational_preference carrier relation" by (simp add: rational_preference_axioms) text \ Multiple convexity alternate definitions intro/dest rules. \ lemma weak_convex1D [dest]: assumes "weak_convex_pref relation" and "x \[relation] y" and "0 \ u" and "0 \ v" and "u + v = 1" shows "u *\<^sub>R x + v *\<^sub>R y \[relation] y" proof- have u_0: "u = 0 \ u *\<^sub>R x + v *\<^sub>R y \[relation] y" proof assume u_0: "u = 0" have "v = 1" using assms(5) u_0 by auto then have "?thesis" by (metis add.left_neutral assms(2) preference.reflexivity preference_axioms real_vector.scale_zero_left refl_onD2 scaleR_one strict_not_refl_weak) thus "u *\<^sub>R x + v *\<^sub>R y \[relation] y " using u_0 by blast qed have "u \ 0 \ u \ 1 \ u *\<^sub>R x + v *\<^sub>R y \[relation] y" by (metis add_cancel_right_right antisym_conv not_le assms weak_convexD ) then show ?thesis by (metis u_0 assms(2,5) add_cancel_right_right real_vector.scale_zero_left scaleR_one) qed lemma weak_convex1I [intro] : assumes "\x. convex (at_least_as_good x carrier relation)" shows "weak_convex_pref relation" proof (rule weak_convexI) fix x and y and \ \ :: real assume assum : "x \[relation] y" assume reals: "0 < \" "0 < \" " \ + \ = 1" then have "x \ carrier" by (meson assum preference.not_outside rational_preference.axioms(1) have_rpr) have "y \ carrier" by (meson assum refl_onD2 reflexivity) then have y_in_upper_cont: "y \ (at_least_as_good y carrier relation)" using assms rational_preference.at_lst_asgd_not_ge rational_preference.compl by (metis empty_iff have_rpr) then have "x \ (at_least_as_good y carrier relation)" using assum pref_in_at_least_as by blast moreover have "0 \ \" and "0 \ \" using reals by (auto) ultimately show "(\ *\<^sub>R x + \ *\<^sub>R y) \[relation] y" by (meson assms(1) at_least_as_goodD convexD reals(3) y_in_upper_cont) qed text \ Definition of convexity in "Handbook of Social Choice and Welfare"@{cite "arrow2010handbook"}. \ lemma convex_def_alt: assumes "rational_preference carrier relation" assumes "weak_convex_pref relation" shows "(\x \ carrier. convex (at_least_as_good x carrier relation))" proof fix x assume x_in: "x \ carrier" show "convex (at_least_as_good x carrier relation)" (is "convex ?x") proof (rule convexI) fix a b :: "'a" and \ :: "real" and \ :: "real" assume a_in: "a \ ?x" assume b_in: "b \ ?x" assume reals: " 0 \ \" "0 \ \" "\ + \ = 1" have a_g_x: "a \[relation] x" using a_in by blast have b_g_x: "b \[relation] x" using b_in by blast have "a \[relation] b \ b \[relation] a" by (meson a_in at_least_as_goodD b_in preference.not_outside rational_preference.compl rational_preference_def assms(1)) then show "\ *\<^sub>R a + \ *\<^sub>R b \ ?x" proof(rule disjE) assume "a \[relation] b" then have "\ *\<^sub>R a + \ *\<^sub>R b \[relation] b" using assms reals by blast then have "\ *\<^sub>R a + \ *\<^sub>R b \[relation] x" by (meson b_g_x assms(1) preference.not_outside x_in rational_preference.strict_is_neg_transitive rational_preference.strict_not_refl_weak rational_preference_def) then show ?thesis by (metis (no_types, lifting) CollectI assms(1) at_least_as_good_def preference_def rational_preference_def) next assume as: "b \[relation] a" then have "\ *\<^sub>R a + \ *\<^sub>R b \[relation] a" by (metis add.commute assms(2) reals weak_convex1D) have "\ *\<^sub>R a + \ *\<^sub>R b \[relation] a" by (metis as add.commute assms(2) reals(1,2,3) weak_convex1D) then have "\ *\<^sub>R a + \ *\<^sub>R b \[relation] x" by (meson a_g_x assms(1) preference.indiff_trans x_in preference.not_outside rational_preference.axioms(1) rational_preference.strict_is_neg_transitive ) then show ?thesis using pref_in_at_least_as by blast qed qed qed lemma convex_imp_convex_str_upper_cnt: assumes "\x \ carrier. convex (at_least_as_good x carrier relation)" shows "convex (at_least_as_good x carrier relation - as_good_as x carrier relation)" (is "convex ( ?a - ?b)") proof (rule convexI) fix a y u v assume as_a: "a \ ?a - ?b" assume as_y: "y \ ?a - ?b" assume reals: "0 \ (u::real)" "0 \ v" "u + v = 1" have cvx: "weak_convex_pref relation " by (meson assms at_least_as_goodD convexI have_rpr preference_def rational_preference.axioms(1) weak_convex1I) then have a_g_x: "a \[relation] x" using as_a by blast then have y_gt_x: "y \[relation] x" using as_y by blast show "u *\<^sub>R a + v *\<^sub>R y \ ?a - ?b" proof show "u *\<^sub>R a + v *\<^sub>R y \ ?a" by (metis DiffD1 a_g_x as_a as_y assms convexD reals have_rpr preference_def rational_preference.axioms(1)) next have "a \[relation] y \ y \[relation] a" by (meson a_g_x y_gt_x assms(1) preference.not_outside have_rpr rational_preference.axioms(1) rational_preference.strict_not_refl_weak) then show "u *\<^sub>R a + v *\<^sub>R y \ ?b" proof assume "a \[relation] y" then have "u *\<^sub>R a + v *\<^sub>R y \[relation] y" using cvx assms(1) reals by blast then have "u *\<^sub>R a + v *\<^sub>R y \[relation] x" using y_gt_x by (meson assms(1) rational_preference.axioms(1) have_rpr rational_preference.strict_is_neg_transitive preference_def) then show "u *\<^sub>R a + v *\<^sub>R y \ as_good_as x carrier relation" by blast next assume "y \[relation] a" then have "u *\<^sub>R a + v *\<^sub>R y \[relation] a" using cvx assms(1) reals by (metis add.commute weak_convex1D) then have "u *\<^sub>R a + v *\<^sub>R y \[relation] x" by (meson a_g_x assms(1) rational_preference.strict_is_neg_transitive rational_preference.axioms(1) preference_def have_rpr) then show "u *\<^sub>R a + v *\<^sub>R y \ ?b" by blast qed qed qed end subsubsection \ Monotone preferences \ definition weak_monotone_prefs :: "'a set \ ('a::ord) relation \ bool" where "weak_monotone_prefs B P \ (\x \ B. \y \ B. x \ y \ x \[P]y)" definition monotone_preference :: "'a set \ ('a::ord) relation \ bool" where "monotone_preference B P \ (\x \ B. \y \ B. x > y \ x \[P] y)" text \ Given a carrier set that is unbounded above (not the "standard" mathematical definition), monotonicity implies local non-satiation. \ lemma unbounded_above_mono_imp_lns: assumes "\M \ carrier. (\x > M. x \ carrier)" assumes mono: "monotone_preference (carrier:: 'a::ordered_euclidean_space set) relation" shows "local_nonsatiation carrier relation" proof(rule lns_distI) fix x and e::"real" assume x_in: "x \ carrier" assume gz : "e > 0" show "\y\carrier. dist y x \ e \ y \[relation] x \ (x, y) \ relation" proof- obtain v :: real where v:"v < e" "0 < v" using gz dense by blast obtain i where i:"(i::'a) \ Basis" by fastforce define y where y_value : "y = x + v *\<^sub>R i" have ge:"y \ x" using y_value i unfolding y_value by (simp add: v(2) zero_le_scaleR_iff) have "y \ x" using y_value i unfolding y_value using v(2) by auto hence y_str_g_x : "y > x" using ge by auto have y_in: "y \ carrier" using assms(1) x_in y_str_g_x by blast then have y_pref_x : "y \[relation] x" using y_str_g_x x_in mono monotone_preference_def by blast hence " norm (y - x) \ e" using \0 < v\ y_value y_value i v by auto hence dist_less_e : "dist y x \ e" by (simp add: dist_norm) thus ?thesis using y_pref_x dist_less_e y_in by blast qed qed end \ No newline at end of file diff --git a/thys/Kuratowski_Closure_Complement/KuratowskiClosureComplementTheorem.thy b/thys/Kuratowski_Closure_Complement/KuratowskiClosureComplementTheorem.thy --- a/thys/Kuratowski_Closure_Complement/KuratowskiClosureComplementTheorem.thy +++ b/thys/Kuratowski_Closure_Complement/KuratowskiClosureComplementTheorem.thy @@ -1,1172 +1,1172 @@ (*<*) theory KuratowskiClosureComplementTheorem imports - "HOL-Analysis.Ordered_Euclidean_Space" + "HOL-Analysis.Multivariate_Analysis" "HOL-Analysis.Continuum_Not_Denumerable" begin (*>*) section\ Introduction \ text\ We discuss a topological curiosity discovered by @{cite "Kuratowski:1922"}: the fact that the number of distinct operators on a topological space generated by compositions of closure and complement never exceeds 14, and is exactly 14 in the case of @{term "\"}. In addition, we prove a theorem due to @{cite "Chagrov:1982"} that classifies topological spaces according to the number of such operators they support. Kuratowski's result, which is exposited in @{cite "TheoremOfTheDay:2015"} and Chapter 7 of @{cite "Chamberland:2015"}, has already been treated in Mizar --- see @{cite "BaginskaGrabowski:2003"} and @{cite "Grabowski:2004"}. To the best of our knowledge, we are the first to mechanize Chagrov's result. Our work is based on a presentation of Kuratowski's and Chagrov's results by @{cite "GardnerJackson:2008"}. We begin with some preliminary facts pertaining to the relationship between interiors of unions and unions of interiors (\S\ref{sec:interiors-unions}) and the relationship between @{term "\"} and @{term "\"} (\S\ref{sec:rat-real}). We then prove Kuratowski's result (\S\ref{sec:kuratowski}) and the corollary that at most 7 distinct operators on a topological space can be generated by compositions of closure and interior (\S\ref{sec:kuratowski-corollary}). Finally, we prove Chagrov's result (\S\ref{sec:chagrov}). \ section\ Interiors and unions \label{sec:interiors-unions} \ definition boundary :: "'a::topological_space set \ 'a set" where "boundary X = closure X - interior X" lemma boundary_empty: shows "boundary {} = {}" unfolding boundary_def by simp definition exterior :: "'a::topological_space set \ 'a set" where "exterior X = - (interior X \ boundary X)" lemma interior_union_boundary: shows "interior (X \ Y) = interior X \ interior Y \ boundary X \ boundary Y \ boundary (X \ Y)" (is "(?lhs1 = ?lhs2) \ ?rhs") proof(rule iffI[OF _ subset_antisym[OF subsetI]]) assume "?lhs1 = ?lhs2" then show ?rhs by (force simp: boundary_def) next fix x assume ?rhs and "x \ ?lhs1" have "x \ ?lhs2" if "x \ interior X" "x \ interior Y" proof(cases "x \ boundary X \ boundary Y") case True with \?rhs\ \x \ ?lhs1\ show ?thesis by (simp add: boundary_def subset_iff) next case False then consider (X) "x \ boundary X" | (Y) "x \ boundary Y" by blast then show ?thesis proof cases case X from X \x \ interior X\ have "x \ exterior X" by (simp add: exterior_def) from \x \ boundary X\ \x \ exterior X\ \x \ interior X\ obtain U where "open U" "U \ - X" "x \ U" by (metis ComplI DiffI boundary_def closure_interior interior_subset open_interior) from \x \ interior (X \ Y)\ obtain U' where "open U'" "U' \ X \ Y" "x \ U'" by (meson interiorE) from \U \ - X\ \U' \ X \ Y\ have "U \ U' \ Y" by blast with \x \ interior Y\ \open U'\ \open U\ \x \ U'\ \x \ U\ show ?thesis by (meson IntI interiorI open_Int) next case Y from Y \x \ interior Y\ have "x \ exterior Y" by (simp add: exterior_def) from \x \ boundary Y\ \x \ exterior Y\ \x \ interior Y\ obtain U where "open U" "U \ - Y" "x \ U" by (metis ComplI DiffI boundary_def closure_interior interior_subset open_interior) from \x \ interior (X \ Y)\ obtain U' where "open U'" "U' \ X \ Y" "x \ U'" by (meson interiorE) from \U \ -Y\ \U' \ X \ Y\ have "U \ U' \ X" by blast with \x \ interior X\ \open U'\ \open U\ \x \ U'\ \x \ U\ show ?thesis by (meson IntI interiorI open_Int) qed qed with \x \ ?lhs1\ show "x \ ?lhs2" by blast next show "?lhs2 \ ?lhs1" by (simp add: interior_mono) qed lemma interior_union_closed_intervals: fixes a :: "'a::ordered_euclidean_space" assumes "b < c" shows "interior ({a..b} \ {c..d}) = interior {a..b} \ interior {c..d}" using assms by (subst interior_union_boundary; auto simp: boundary_def) section\ Additional facts about the rationals and reals \label{sec:rat-real} \ lemma Rat_real_limpt: fixes x :: real shows "x islimpt \" proof(rule islimptI) fix T assume "x \ T" "open T" then obtain e where "0 < e" and ball: "\x' - x\ < e \ x' \ T" for x' by (auto simp: open_real) from \0 < e\ obtain q where "x < real_of_rat q \ real_of_rat q < x + e" using of_rat_dense by force with ball show "\y\\. y \ T \ y \ x" by force qed lemma Rat_closure: shows "closure \ = (UNIV :: real set)" unfolding closure_def using Rat_real_limpt by blast lemma Rat_interval_closure: fixes x :: real assumes "x < y" shows "closure ({x<.. \) = {x..y}" using assms by (metis (no_types, lifting) Rat_closure closure_closure closure_greaterThanLessThan closure_mono inf_le1 inf_top.right_neutral open_Int_closure_subset open_real_greaterThanLessThan subset_antisym) lemma Rat_not_open: fixes T :: "real set" assumes "open T" assumes "T \ {}" shows "\T \ \" using assms by (simp add: countable_rat open_minus_countable subset_eq) lemma Irrat_dense_in_real: fixes x :: real assumes "x < y" shows "\r\-\. x < r \ r < y" using assms Rat_not_open[where T="{x<.. - {y..z} = - {x..z}" using assms by auto section\ Kuratowski's result \label{sec:kuratowski} \ text\ We prove that at most 14 distinct operators can be generated by compositions of @{const "closure"} and complement. For convenience, we give these operators short names and try to avoid pointwise reasoning. We treat the @{const "interior"} operator at the same time. \ declare o_apply[simp del] definition C :: "'a::topological_space set \ 'a set" where "C X = - X" definition K :: "'a::topological_space set \ 'a set" where "K X = closure X" definition I :: "'a::topological_space set \ 'a set" where "I X = interior X" lemma C_C: shows "C \ C = id" by (simp add: fun_eq_iff C_def o_apply) lemma K_K: shows "K \ K = K" by (simp add: fun_eq_iff K_def o_apply) lemma I_I: shows "I \ I = I" unfolding I_def by (simp add: o_def) lemma I_K: shows "I = C \ K \ C" unfolding C_def I_def K_def by (simp add: o_def interior_closure) lemma K_I: shows "K = C \ I \ C" unfolding C_def I_def K_def by (simp add: o_def interior_closure) lemma K_I_K_I: shows "K \ I \ K \ I = K \ I" unfolding C_def I_def K_def by (clarsimp simp: fun_eq_iff o_apply closure_minimal closure_mono closure_subset interior_maximal interior_subset subset_antisym) lemma I_K_I_K: shows "I \ K \ I \ K = I \ K" unfolding C_def I_def K_def by (simp add: fun_eq_iff o_apply) (metis (no_types) closure_closure closure_mono closure_subset interior_maximal interior_mono interior_subset open_interior subset_antisym) lemma K_mono: assumes "x \ y" shows "K x \ K y" using assms unfolding K_def by (simp add: closure_mono) text\ The following lemma embodies the crucial observation about compositions of @{const "C"} and @{const "K"}: \ lemma KCKCKCK_KCK: shows "K \ C \ K \ C \ K \ C \ K = K \ C \ K" (is "?lhs = ?rhs") proof(rule ext[OF equalityI]) fix x have "(C \ K \ C \ K \ C \ K) x \ ?rhs x" by (simp add: C_def K_def closure_def o_apply) then have "(K \ (C \ K \ C \ K \ C \ K)) x \ (K \ ?rhs) x" by (simp add: K_mono o_apply) then show "?lhs x \ ?rhs x" by (simp add: K_K o_assoc) next fix x :: "'a::topological_space set" have "(C \ K \ C \ K) x \ K x" by (simp add: C_def K_def closure_def o_apply) then have "(K \ (C \ K \ C \ K)) x \ (K \ K) x" by (simp add: K_mono o_apply) then have "(C \ (K \ K)) x \ (C \ (K \ (C \ K \ C \ K))) x" by (simp add: C_def o_apply) then have "(K \ (C \ (K \ K))) x \ (K \ (C \ (K \ (C \ K \ C \ K)))) x" by (simp add: K_mono o_apply) then show "?rhs x \ ?lhs x" by (simp add: K_K o_assoc) qed text\ The inductive set \CK\ captures all operators that can be generated by compositions of @{const "C"} and @{const "K"}. We shallowly embed the operators; that is, we identify operators up to extensional equality. \ inductive CK :: "('a::topological_space set \ 'a set) \ bool" where "CK C" | "CK K" | "\ CK f; CK g \ \ CK (f \ g)" declare CK.intros[intro!] lemma CK_id[intro!]: "CK id" by (metis CK.intros(1) CK.intros(3) C_C) text\ The inductive set \CK_nf\ captures the normal forms for the 14 distinct operators. \ inductive CK_nf :: "('a::topological_space set \ 'a set) \ bool" where "CK_nf id" | "CK_nf C" | "CK_nf K" | "CK_nf (C \ K)" | "CK_nf (K \ C)" | "CK_nf (C \ K \ C)" | "CK_nf (K \ C \ K)" | "CK_nf (C \ K \ C \ K)" | "CK_nf (K \ C \ K \ C)" | "CK_nf (C \ K \ C \ K \ C)" | "CK_nf (K \ C \ K \ C \ K)" | "CK_nf (C \ K \ C \ K \ C \ K)" | "CK_nf (K \ C \ K \ C \ K \ C)" | "CK_nf (C \ K \ C \ K \ C \ K \ C)" declare CK_nf.intros[intro!] lemma CK_nf_set: shows "{f . CK_nf f} = {id, C, K, C \ K, K \ C, C \ K \ C, K \ C \ K, C \ K \ C \ K, K \ C \ K \ C, C \ K \ C \ K \ C, K \ C \ K \ C \ K, C \ K \ C \ K \ C \ K, K \ C \ K \ C \ K \ C, C \ K \ C \ K \ C \ K \ C}" by (auto simp: CK_nf.simps) text\ That each operator generated by compositions of @{const "C"} and @{const "K"} is extensionally equivalent to one of the normal forms captured by \CK_nf\ is demonstrated by means of an induction over the construction of \CK_nf\ and an appeal to the facts proved above. \ theorem CK_nf: "CK f \ CK_nf f" proof(rule iffI) assume "CK f" then show "CK_nf f" by induct (elim CK_nf.cases; clarsimp simp: id_def[symmetric] C_C K_K KCKCKCK_KCK o_assoc; simp add: o_assoc[symmetric]; clarsimp simp: C_C K_K KCKCKCK_KCK o_assoc | blast)+ next assume "CK_nf f" then show "CK f" by induct (auto simp: id_def[symmetric]) qed theorem CK_card: shows "card {f. CK f} \ 14" by (auto simp: CK_nf CK_nf_set card_insert intro!: le_trans[OF card_Diff1_le]) text\ We show, using the following subset of \\\ (an example taken from @{cite "Rusin:2001"}) as a witness, that there exist topological spaces on which all 14 operators are distinct. \ definition RRR :: "real set" where "RRR = {0<..<1} \ {1<..<2} \ {3} \ ({5<..<7} \ \)" text\ The following facts allow the required proofs to proceed by @{method simp}: \ lemma RRR_closure: shows "closure RRR = {0..2} \ {3} \ {5..7}" unfolding RRR_def by (force simp: closure_insert Rat_interval_closure) lemma RRR_interior: "interior RRR = {0<..<1} \ {1<..<2}" (is "?lhs = ?rhs") proof(rule equalityI[OF subsetI subsetI]) fix x assume "x \ ?lhs" then obtain T where "open T" and "x \ T" and "T \ RRR" by (blast elim: interiorE) then obtain e where "0 < e" and "ball x e \ T" by (blast elim!: openE) from \x \ T\ \0 < e\ \ball x e \ T\ \T \ RRR\ have False if "x = 3" using that unfolding RRR_def ball_def by (auto dest!: subsetD[where c="min (3 + e/2) 4"] simp: dist_real_def) moreover from Irrat_dense_in_real[where x="x" and y="x + e/2"] \0 < e\ obtain r where "r\- \ \ x < r \ r < x + e / 2" by auto with \x \ T\ \ball x e \ T\ \T \ RRR\ have False if "x \ {5<..<7} \ \" using that unfolding RRR_def ball_def by (force simp: dist_real_def dest: subsetD[where c=r]) moreover note \x \ interior RRR\ ultimately show "x \ ?rhs" unfolding RRR_def by (auto dest: subsetD[OF interior_subset]) next fix x assume "x \ ?rhs" then show "x \ ?lhs" unfolding RRR_def interior_def by (auto intro: open_real_greaterThanLessThan) qed lemma RRR_interior_closure[simplified]: shows "interior ({0::real..2} \ {3} \ {5..7}) = {0<..<2} \ {5<..<7}" (is "?lhs = ?rhs") proof - have "?lhs = interior ({0..2} \ {5..7})" by (metis (no_types, lifting) Un_assoc Un_commute closed_Un closed_eucl_atLeastAtMost interior_closed_Un_empty_interior interior_singleton) also have "... = ?rhs" by (simp add: interior_union_closed_intervals) finally show ?thesis . qed text\ The operators can be distinguished by testing which of the points in \{1,2,3,4,6}\ belong to their results. \ definition test :: "(real set \ real set) \ bool list" where "test f \ map (\x. x \ f RRR) [1,2,3,4,6]" lemma RRR_test: assumes "f RRR = g RRR" shows "test f = test g" unfolding test_def using assms by simp lemma nf_RRR: shows "test id = [False, False, True, False, True]" "test C = [True, True, False, True, False]" "test K = [True, True, True, False, True]" "test (K \ C) = [True, True, True, True, True]" "test (C \ K) = [False, False, False, True, False]" "test (C \ K \ C) = [False, False, False, False, False]" "test (K \ C \ K) = [False, True, True, True, False]" "test (C \ K \ C \ K) = [True, False, False, False, True]" "test (K \ C \ K \ C) = [True, True, False, False, False]" "test (C \ K \ C \ K \ C) = [False, False, True, True, True]" "test (K \ C \ K \ C \ K) = [True, True, False, False, True]" "test (C \ K \ C \ K \ C \ K) = [False, False, True, True, False]" "test (K \ C \ K \ C \ K \ C) = [False, True, True, True, True]" "test (C \ K \ C \ K \ C \ K \ C) = [True, False, False, False, False]" unfolding test_def C_def K_def by (simp_all add: RRR_closure RRR_interior RRR_interior_closure closure_complement closed_interval_Int_compl o_apply) (simp_all add: RRR_def) theorem CK_nf_real_card: shows "card ((\ f. f RRR) ` {f . CK_nf f}) = 14" by (simp add: CK_nf_set) ((subst card_insert_disjoint; auto dest!: RRR_test simp: nf_RRR id_def[symmetric])[1])+ theorem CK_real_card: shows "card {f::real set \ real set. CK f} = 14" (is "?lhs = ?rhs") proof(rule antisym[OF CK_card]) show "?rhs \ ?lhs" unfolding CK_nf by (rule le_trans[OF eq_imp_le[OF CK_nf_real_card[symmetric]] card_image_le]) (simp add: CK_nf_set) qed section\ A corollary of Kuratowski's result \label{sec:kuratowski-corollary} \ text\ We show that it is a corollary of @{thm [source] CK_real_card} that at most 7 distinct operators on a topological space can be generated by compositions of closure and interior. In the case of \\\, exactly 7 distinct operators can be so generated. \ inductive IK :: "('a::topological_space set \ 'a set) \ bool" where "IK id" | "IK I" | "IK K" | "\ IK f; IK g \ \ IK (f \ g)" inductive IK_nf :: "('a::topological_space set \ 'a set) \ bool" where "IK_nf id" | "IK_nf I" | "IK_nf K" | "IK_nf (I \ K)" | "IK_nf (K \ I)" | "IK_nf (I \ K \ I)" | "IK_nf (K \ I \ K)" declare IK.intros[intro!] declare IK_nf.intros[intro!] lemma IK_nf_set: "{f . IK_nf f} = {id, I, K, I \ K, K \ I, I \ K \ I, K \ I \ K}" by (auto simp: IK_nf.simps) theorem IK_nf: "IK f \ IK_nf f" proof(rule iffI) assume "IK f" then show "IK_nf f" by induct (elim IK_nf.cases; clarsimp simp: id_def[symmetric] o_assoc; simp add: I_I K_K o_assoc[symmetric]; clarsimp simp: K_I_K_I I_K_I_K o_assoc | blast)+ next assume "IK_nf f" then show "IK f" by induct blast+ qed theorem IK_card: shows "card {f. IK f} \ 7" by (auto simp: IK_nf IK_nf_set card_insert intro!: le_trans[OF card_Diff1_le]) theorem IK_nf_real_card: shows "card ((\ f. f RRR) ` {f . IK_nf f}) = 7" by (simp add: IK_nf_set) ((subst card_insert_disjoint; auto dest!: RRR_test simp: nf_RRR I_K id_def[symmetric] o_assoc)[1])+ theorem IK_real_card: shows "card {f::real set \ real set. IK f} = 7" (is "?lhs = ?rhs") proof(rule antisym[OF IK_card]) show "?rhs \ ?lhs" unfolding IK_nf by (rule le_trans[OF eq_refl[OF IK_nf_real_card[symmetric]] card_image_le]) (simp add: IK_nf_set) qed section\ Chagrov's result \label{sec:chagrov} \ text\ Chagrov's theorem, which is discussed in Section 2.1 of @{cite "GardnerJackson:2008"}, states that the number of distinct operators on a topological space that can be generated by compositions of closure and complement is one of 2, 6, 8, 10 or 14. We begin by observing that the set of normal forms @{const "CK_nf"} can be split into two disjoint sets, @{term "CK_nf_pos"} and @{term "CK_nf_neg"}, which we define in terms of interior and closure. \ inductive CK_nf_pos :: "('a::topological_space set \ 'a set) \ bool" where "CK_nf_pos id" | "CK_nf_pos I" | "CK_nf_pos K" | "CK_nf_pos (I \ K)" | "CK_nf_pos (K \ I)" | "CK_nf_pos (I \ K \ I)" | "CK_nf_pos (K \ I \ K)" declare CK_nf_pos.intros[intro!] lemma CK_nf_pos_set: shows "{f . CK_nf_pos f} = {id, I, K, I \ K, K \ I, I \ K \ I, K \ I \ K}" by (auto simp: CK_nf_pos.simps) definition CK_nf_neg :: "('a::topological_space set \ 'a set) \ bool" where "CK_nf_neg f \ (\g. CK_nf_pos g \ f = C \ g)" lemma CK_nf_pos_neg_disjoint: assumes "CK_nf_pos f" assumes "CK_nf_neg g" shows "f \ g" using assms unfolding CK_nf_neg_def by (clarsimp simp: CK_nf_pos.simps; elim disjE; metis comp_def C_def I_def K_def Compl_iff closure_UNIV interior_UNIV id_apply) lemma CK_nf_pos_neg_CK_nf: "CK_nf f \ CK_nf_pos f \ CK_nf_neg f" (is "?lhs \ ?rhs") proof(rule iffI) assume ?lhs then show ?rhs unfolding CK_nf_neg_def by (rule CK_nf.cases; metis (no_types, lifting) CK_nf_pos.simps C_C I_K K_I comp_id o_assoc) next assume ?rhs then show ?lhs unfolding CK_nf_neg_def by (auto elim!: CK_nf_pos.cases simp: I_K C_C o_assoc) qed text\ We now focus on @{const "CK_nf_pos"}. In particular, we show that its cardinality for any given topological space is one of 1, 3, 4, 5 or 7. The proof consists of exhibiting normal forms for the operators supported by each of six classes of topological spaces. These are sublattices of the following lattice of @{const CK_nf_pos} operators: \ lemmas K_I_K_subseteq_K = closure_mono[OF interior_subset, of "closure X", simplified] for X lemma CK_nf_pos_lattice: shows "I \ (id :: 'a::topological_space set \ 'a set)" "id \ (K :: 'a::topological_space set \ 'a set)" "I \ I \ K \ (I :: 'a::topological_space set \ 'a set)" "I \ K \ I \ I \ (K :: 'a::topological_space set \ 'a set)" "I \ K \ I \ K \ (I :: 'a::topological_space set \ 'a set)" "I \ K \ K \ I \ (K :: 'a::topological_space set \ 'a set)" "K \ I \ K \ I \ (K :: 'a::topological_space set \ 'a set)" "K \ I \ K \ (K :: 'a::topological_space set \ 'a set)" unfolding I_def K_def by (simp_all add: interior_subset closure_subset interior_maximal closure_mono o_apply interior_mono K_I_K_subseteq_K le_funI) text\ We define the six classes of topological spaces in question, and show that they are related by inclusion in the following way (as shown in Figure 2.3 of @{cite "GardnerJackson:2008"}): \begin{center} \begin{tikzpicture}[scale=.7] \node (d) at (30:5cm) {Open unresolvable spaces}; \node (one) at (90:5cm) {Kuratowski spaces}; \node (b) at (150:5cm) {Extremally disconnected spaces}; \node (a) at (210:5cm) {Partition spaces}; \node (zero) at (270:5cm) {Discrete spaces}; \node (c) at (330:5cm) {\hspace{1cm}Extremally disconnected and open unresolvable spaces}; \draw (zero) -- (a) -- (b) -- (one) -- (d) -- (c) -- (zero); \draw (c) -- (b); \end{tikzpicture} \end{center} \ subsection\ Discrete spaces \ definition "discrete (X :: 'a::topological_space set) \ I = (id::'a set \ 'a set)" lemma discrete_eqs: assumes "discrete (X :: 'a::topological_space set)" shows "I = (id::'a set \ 'a set)" "K = (id::'a set \ 'a set)" using assms unfolding discrete_def by (auto simp: C_C K_I) lemma discrete_card: assumes "discrete (X :: 'a::topological_space set)" shows "card {f. CK_nf_pos (f::'a set \ 'a set)} = 1" using discrete_eqs[OF assms] CK_nf_pos_lattice[where 'a='a] by (simp add: CK_nf_pos_set) lemma discrete_discrete_topology: fixes X :: "'a::topological_space set" assumes "\Y::'a set. open Y" shows "discrete X" using assms unfolding discrete_def I_def interior_def islimpt_def by (auto simp: fun_eq_iff) subsection\ Partition spaces \ definition "part (X :: 'a::topological_space set) \ K \ I = (I :: 'a set \ 'a set)" lemma discrete_part: assumes "discrete X" shows "part X" using assms unfolding discrete_def part_def by (simp add: C_C K_I) lemma part_eqs: assumes "part (X :: 'a::topological_space set)" shows "K \ I = (I :: 'a set \ 'a set)" "I \ K = (K :: 'a set \ 'a set)" using assms unfolding part_def by (assumption, metis (no_types, hide_lams) I_I K_I o_assoc) lemma part_not_discrete_card: assumes "part (X :: 'a::topological_space set)" assumes "\discrete X" shows "card {f. CK_nf_pos (f::'a set \ 'a set)} = 3" using part_eqs[OF \part X\] \\discrete X\ CK_nf_pos_lattice[where 'a='a] unfolding discrete_def by (simp add: CK_nf_pos_set card_insert_if C_C I_K K_K o_assoc; metis comp_id) text\ A partition space is a topological space whose basis consists of the empty set and the equivalence classes of points of the space induced by some equivalence relation @{term "R"} on the underlying set of the space. Equivalently, a partition space is one in which every open set is closed. Thus, for example, the class of partition spaces includes every topological space whose open sets form a boolean algebra. \ datatype part_witness = a | b | c lemma part_witness_UNIV: shows "UNIV = set [a, b, c]" using part_witness.exhaust by auto lemmas part_witness_pow = subset_subseqs[OF subset_trans[OF subset_UNIV Set.equalityD1[OF part_witness_UNIV]]] lemmas part_witness_Compl = Compl_eq_Diff_UNIV[where 'a=part_witness, unfolded part_witness_UNIV, simplified] instantiation part_witness :: topological_space begin definition "open_part_witness X \ X \ {{}, {a}, {b, c}, {a, b, c}}" lemma part_witness_ball: "(\s\S. s \ {{}, {a}, {b, c}, {a, b, c}}) \ S \ set [{}, {a}, {b, c}, {a, b, c}]" by simp blast lemmas part_witness_subsets_pow = subset_subseqs[OF iffD1[OF part_witness_ball]] instance proof standard fix K :: "part_witness set set" assume "\S\K. open S" then show "open (\K)" unfolding open_part_witness_def by - (drule part_witness_subsets_pow; clarsimp; elim disjE; simp add: insert_commute) qed (auto simp: open_part_witness_def part_witness_UNIV) end lemma part_witness_interior_simps: shows "interior {a} = {a}" "interior {b} = {}" "interior {c} = {}" "interior {a, b} = {a}" "interior {a, c} = {a}" "interior {b, c} = {b, c}" "interior {a, b, c} = {a, b, c}" unfolding interior_def open_part_witness_def by auto lemma part_witness_part: fixes X :: "part_witness set" shows "part X" proof - have "closure (interior Y) = interior Y" for Y :: "part_witness set" using part_witness_pow[where X=Y] by (auto simp: closure_interior part_witness_interior_simps part_witness_Compl insert_Diff_if) then show ?thesis unfolding part_def I_def K_def by (simp add: o_def) qed lemma part_witness_not_discrete: fixes X :: "part_witness set" shows "\discrete X" unfolding discrete_def I_def by (clarsimp simp: o_apply fun_eq_iff exI[where x="{b}"] part_witness_interior_simps) lemma part_witness_card: shows "card {f. CK_nf_pos (f::part_witness set \ part_witness set)} = 3" by (rule part_not_discrete_card[OF part_witness_part part_witness_not_discrete]) subsection\ Extremally disconnected and open unresolvable spaces \ definition "ed_ou (X :: 'a::topological_space set) \ I \ K = K \ (I :: 'a set \ 'a set)" lemma discrete_ed_ou: assumes "discrete X" shows "ed_ou X" using assms unfolding discrete_def ed_ou_def by simp lemma ed_ou_eqs: assumes "ed_ou (X :: 'a::topological_space set)" shows "I \ K \ I = K \ (I :: 'a set \ 'a set)" "K \ I \ K = K \ (I :: 'a set \ 'a set)" "I \ K = K \ (I :: 'a set \ 'a set)" using assms unfolding ed_ou_def by (metis I_I K_K o_assoc)+ lemma ed_ou_neqs: assumes "ed_ou (X :: 'a::topological_space set)" assumes "\discrete X" shows "I \ (K :: 'a set \ 'a set)" "I \ K \ (I :: 'a set \ 'a set)" "K \ K \ (I :: 'a set \ 'a set)" "I \ (id :: 'a set \ 'a set)" "K \ (id :: 'a set \ 'a set)" using assms CK_nf_pos_lattice[where 'a='a] unfolding ed_ou_def discrete_def by (metis (no_types, lifting) C_C I_K K_I comp_id o_assoc antisym)+ lemma ed_ou_not_discrete_card: assumes "ed_ou (X :: 'a::topological_space set)" assumes "\discrete X" shows "card {f. CK_nf_pos (f::'a set \ 'a set)} = 4" using ed_ou_eqs[OF \ed_ou X\] ed_ou_neqs[OF assms] by (subst CK_nf_pos_set) (subst card_insert_disjoint; (auto)[1])+ text\ We consider an example extremally disconnected and open unresolvable topological space. \ datatype ed_ou_witness = a | b | c | d | e lemma ed_ou_witness_UNIV: shows "UNIV = set [a, b, c, d, e]" using ed_ou_witness.exhaust by auto lemmas ed_ou_witness_pow = subset_subseqs[OF subset_trans[OF subset_UNIV Set.equalityD1[OF ed_ou_witness_UNIV]]] lemmas ed_ou_witness_Compl = Compl_eq_Diff_UNIV[where 'a=ed_ou_witness, unfolded ed_ou_witness_UNIV, simplified] instance ed_ou_witness :: finite by standard (simp add: ed_ou_witness_UNIV) instantiation ed_ou_witness :: topological_space begin inductive open_ed_ou_witness :: "ed_ou_witness set \ bool" where "open_ed_ou_witness {}" | "open_ed_ou_witness {a}" | "open_ed_ou_witness {b}" | "open_ed_ou_witness {e}" | "open_ed_ou_witness {a, c}" | "open_ed_ou_witness {b, d}" | "open_ed_ou_witness {a, c, e}" | "open_ed_ou_witness {a, b}" | "open_ed_ou_witness {a, e}" | "open_ed_ou_witness {b, e}" | "open_ed_ou_witness {a, b, c}" | "open_ed_ou_witness {a, b, d}" | "open_ed_ou_witness {a, b, e}" | "open_ed_ou_witness {b, d, e}" | "open_ed_ou_witness {a, b, c, d}" | "open_ed_ou_witness {a, b, c, e}" | "open_ed_ou_witness {a, b, d, e}" | "open_ed_ou_witness {a, b, c, d, e}" declare open_ed_ou_witness.intros[intro!] lemma ed_ou_witness_inter: fixes S :: "ed_ou_witness set" assumes "open S" assumes "open T" shows "open (S \ T)" using assms by (auto elim!: open_ed_ou_witness.cases) lemma ed_ou_witness_union: fixes X :: "ed_ou_witness set set" assumes "\x\X. open x" shows "open (\X)" using finite[of X] assms by (induct, force) (clarsimp; elim open_ed_ou_witness.cases; simp add: open_ed_ou_witness.simps subset_insertI2 insert_commute; metis Union_empty_conv) instance by standard (auto simp: ed_ou_witness_UNIV intro: ed_ou_witness_inter ed_ou_witness_union) end lemma ed_ou_witness_interior_simps: shows "interior {a} = {a}" "interior {b} = {b}" "interior {c} = {}" "interior {d} = {}" "interior {e} = {e}" "interior {a, b} = {a, b}" "interior {a, c} = {a, c}" "interior {a, d} = {a}" "interior {a, e} = {a, e}" "interior {b, c} = {b}" "interior {b, d} = {b, d}" "interior {b, e} = {b, e}" "interior {c, d} = {}" "interior {c, e} = {e}" "interior {d, e} = {e}" "interior {a, b, c} = {a, b, c}" "interior {a, b, d} = {a, b, d}" "interior {a, b, e} = {a, b, e}" "interior {a, c, d} = {a, c}" "interior {a, c, e} = {a, c, e}" "interior {a, d, e} = {a, e}" "interior {b, c, d} = {b, d}" "interior {b, c, e} = {b, e}" "interior {b, d, e} = {b, d, e}" "interior {c, d, e} = {e}" "interior {a, b, c, d} = {a, b, c, d}" "interior {a, b, c, e} = {a, b, c, e}" "interior {a, b, d, e} = {a, b, d, e}" "interior {a, b, c, d, e} = {a, b, c, d, e}" "interior {a, c, d, e} = {a, c, e}" "interior {b, c, d, e} = {b, d, e}" unfolding interior_def by safe (clarsimp simp: open_ed_ou_witness.simps; blast)+ lemma ed_ou_witness_not_discrete: fixes X :: "ed_ou_witness set" shows "\discrete X" unfolding discrete_def I_def using ed_ou_witness_interior_simps by (force simp: fun_eq_iff) lemma ed_ou_witness_ed_ou: fixes X :: "ed_ou_witness set" shows "ed_ou X" unfolding ed_ou_def I_def K_def proof(clarsimp simp: o_apply fun_eq_iff) fix x :: "ed_ou_witness set" from ed_ou_witness_pow[of x] show "interior (closure x) = closure (interior x)" by - (simp; elim disjE; simp add: closure_interior ed_ou_witness_interior_simps ed_ou_witness_Compl insert_Diff_if) qed lemma ed_ou_witness_card: shows "card {f. CK_nf_pos (f::ed_ou_witness set \ ed_ou_witness set)} = 4" by (rule ed_ou_not_discrete_card[OF ed_ou_witness_ed_ou ed_ou_witness_not_discrete]) subsection\ Extremally disconnected spaces \ definition "extremally_disconnected (X :: 'a::topological_space set) \ K \ I \ K = I \ (K :: 'a set \ 'a set)" lemma ed_ou_part_extremally_disconnected: assumes "ed_ou X" assumes "part X" shows "extremally_disconnected X" using assms unfolding extremally_disconnected_def ed_ou_def part_def by simp lemma extremally_disconnected_eqs: fixes X :: "'a::topological_space set" assumes "extremally_disconnected X" shows "I \ K \ I = K \ (I :: 'a set \ 'a set)" "K \ I \ K = I \ (K :: 'a set \ 'a set)" using assms unfolding extremally_disconnected_def by (metis K_I_K_I)+ lemma extremally_disconnected_not_part_not_ed_ou_card: fixes X :: "'a::topological_space set" assumes "extremally_disconnected X" assumes "\part X" assumes "\ed_ou X" shows "card {f. CK_nf_pos (f::'a set \ 'a set)} = 5" using extremally_disconnected_eqs[OF \extremally_disconnected X\] CK_nf_pos_lattice[where 'a='a] assms(2,3) unfolding part_def ed_ou_def by (simp add: CK_nf_pos_set C_C I_K K_K o_assoc card_insert_if; metis (no_types) C_C K_I id_comp o_assoc) text\ Any topological space having an infinite underlying set and whose topology consists of the empty set and every cofinite subset of the underlying set is extremally disconnected. We consider an example such space having a countably infinite underlying set. \ datatype 'a cofinite = cofinite 'a instantiation cofinite :: (type) topological_space begin definition "open_cofinite = (\X::'a cofinite set. finite (-X) \ X = {})" instance by standard (auto simp: open_cofinite_def uminus_Sup) end lemma cofinite_closure_finite: fixes X :: "'a cofinite set" assumes "finite X" shows "closure X = X" using assms by (simp add: closed_open open_cofinite_def) lemma cofinite_closure_infinite: fixes X :: "'a cofinite set" assumes "infinite X" shows "closure X = UNIV" using assms by (metis Compl_empty_eq closure_subset double_compl finite_subset interior_complement open_cofinite_def open_interior) lemma cofinite_interior_finite: fixes X :: "'a cofinite set" assumes "finite X" assumes "infinite (UNIV::'a cofinite set)" shows "interior X = {}" using assms cofinite_closure_infinite[where X="-X"] by (simp add: interior_closure) lemma cofinite_interior_infinite: fixes X :: "'a cofinite set" assumes "infinite X" assumes "infinite (-X)" shows "interior X = {}" using assms cofinite_closure_infinite[where X="-X"] by (simp add: interior_closure) abbreviation "evens :: nat cofinite set \ {cofinite n | n. \i. n=2*i}" lemma evens_infinite: shows "infinite evens" proof(rule iffD2[OF infinite_iff_countable_subset], rule exI, rule conjI) let ?f = "\n::nat. cofinite (2*n)" show "inj ?f" by (auto intro: inj_onI) show "range ?f \ evens" by auto qed lemma cofinite_nat_infinite: shows "infinite (UNIV::nat cofinite set)" using evens_infinite finite_Diff2 by fastforce lemma evens_Compl_infinite: shows "infinite (- evens)" proof(rule iffD2[OF infinite_iff_countable_subset], rule exI, rule conjI) let ?f = "\n::nat. cofinite (2*n+1)" show "inj ?f" by (auto intro: inj_onI) show "range ?f \ -evens" by clarsimp presburger qed lemma evens_closure: shows "closure evens = UNIV" using evens_infinite by (rule cofinite_closure_infinite) lemma evens_interior: shows "interior evens = {}" using evens_infinite evens_Compl_infinite by (rule cofinite_interior_infinite) lemma cofinite_not_part: fixes X :: "nat cofinite set" shows "\part X" unfolding part_def I_def K_def using cofinite_nat_infinite by (clarsimp simp: fun_eq_iff o_apply) (metis (no_types) cofinite_closure_finite cofinite_interior_finite double_compl finite.emptyI finite.insertI insert_not_empty interior_closure) lemma cofinite_not_ed_ou: fixes X :: "nat cofinite set" shows "\ed_ou X" unfolding ed_ou_def I_def K_def by (clarsimp simp: fun_eq_iff o_apply evens_closure evens_interior exI[where x="evens"]) lemma cofinite_extremally_disconnected_aux: fixes X :: "nat cofinite set" shows "closure (interior (closure X)) \ interior (closure X)" by (metis subsetI closure_closure closure_complement closure_def closure_empty finite_Un interior_eq open_cofinite_def open_interior) lemma cofinite_extremally_disconnected: fixes X :: "nat cofinite set" shows "extremally_disconnected X" unfolding extremally_disconnected_def I_def K_def by (auto simp: fun_eq_iff o_apply dest: subsetD[OF closure_subset] subsetD[OF interior_subset] subsetD[OF cofinite_extremally_disconnected_aux]) lemma cofinite_card: shows "card {f. CK_nf_pos (f::nat cofinite set \ nat cofinite set)} = 5" by (rule extremally_disconnected_not_part_not_ed_ou_card[OF cofinite_extremally_disconnected cofinite_not_part cofinite_not_ed_ou]) subsection\ Open unresolvable spaces \ definition "open_unresolvable (X :: 'a::topological_space set) \ K \ I \ K = K \ (I :: 'a set \ 'a set)" lemma ed_ou_open_unresolvable: assumes "ed_ou X" shows "open_unresolvable X" using assms unfolding open_unresolvable_def by (simp add: ed_ou_eqs) lemma open_unresolvable_eqs: assumes "open_unresolvable (X :: 'a::topological_space set)" shows "I \ K \ I = I \ (K :: 'a set \ 'a set)" "K \ I \ K = K \ (I :: 'a set \ 'a set)" using assms unfolding open_unresolvable_def by - (metis I_K_I_K o_assoc; simp) lemma not_ed_ou_neqs: assumes "\ed_ou (X :: 'a::topological_space set)" shows "I \ I \ (K :: 'a set \ 'a set)" "K \ K \ (I :: 'a set \ 'a set)" using assms unfolding ed_ou_def by (simp_all add: fun_eq_iff I_K K_def C_def o_apply) (metis (no_types, hide_lams) closure_eq_empty disjoint_eq_subset_Compl double_complement interior_Int interior_complement set_eq_subset)+ lemma open_unresolvable_not_ed_ou_card: assumes "open_unresolvable (X :: 'a::topological_space set)" assumes "\ed_ou X" shows "card {f. CK_nf_pos (f::'a set \ 'a set)} = 5" using open_unresolvable_eqs[OF \open_unresolvable X\] not_ed_ou_neqs[OF \\ed_ou X\] \\ed_ou X\ unfolding ed_ou_def by (auto simp: CK_nf_pos_set card_insert_if) text\ We show that the class of open unresolvable spaces is non-empty by exhibiting an example of such a space. \ datatype ou_witness = a | b | c lemma ou_witness_UNIV: shows "UNIV = set [a, b, c]" using ou_witness.exhaust by auto instantiation ou_witness :: topological_space begin definition "open_ou_witness X \ a \ X \ X = UNIV" instance by standard (auto simp: open_ou_witness_def) end lemma ou_witness_closure_simps: shows "closure {a} = {a}" "closure {b} = {a, b}" "closure {c} = {a, c}" "closure {a, b} = {a, b}" "closure {a, c} = {a, c}" "closure {a, b, c} = {a, b, c}" "closure {b, c} = {a, b, c}" unfolding closure_def islimpt_def open_ou_witness_def by force+ lemma ou_witness_open_unresolvable: fixes X :: "ou_witness set" shows "open_unresolvable X" unfolding open_unresolvable_def I_def K_def by (clarsimp simp: o_apply fun_eq_iff) (metis (no_types, lifting) Compl_iff K_I_K_subseteq_K closure_complement closure_interior closure_mono closure_subset interior_eq interior_maximal open_ou_witness_def subset_antisym) lemma ou_witness_not_ed_ou: fixes X :: "ou_witness set" shows "\ed_ou X" unfolding ed_ou_def I_def K_def by (clarsimp simp: o_apply fun_eq_iff) (metis UNIV_I insert_iff interior_eq open_ou_witness_def singletonD ou_witness.distinct(4,5) ou_witness.simps(2) ou_witness_closure_simps(2)) lemma ou_witness_card: shows "card {f. CK_nf_pos (f::ou_witness set \ ou_witness set)} = 5" by (rule open_unresolvable_not_ed_ou_card[OF ou_witness_open_unresolvable ou_witness_not_ed_ou]) subsection\ Kuratowski spaces \ definition "kuratowski (X :: 'a::topological_space set) \ \extremally_disconnected X \ \open_unresolvable X" text\ A Kuratowski space distinguishes all 7 positive operators. \ lemma part_closed_open: fixes X :: "'a::topological_space set" assumes "I \ K \ I = (I::'a set \ 'a set)" assumes "closed X" shows "open X" proof(rule Topological_Spaces.openI) fix x assume "x \ X" let ?S = "I (-{x})" let ?G = "-K ?S" have "x \ ?G" proof - from \I \ K \ I = I\ have "I (K (I ?S)) = ?S" "I ?S = ?S" unfolding I_def K_def by (simp_all add: o_def fun_eq_iff) then have "K (I ?S) \ UNIV" unfolding I_def K_def using interior_subset by fastforce moreover have "G \ ?S \ x \ G" if "open G" for G using that unfolding I_def by (meson interior_maximal subset_Compl_singleton) ultimately show ?thesis unfolding I_def K_def by clarsimp (metis (no_types, lifting) ComplD Compl_empty_eq closure_interior closure_subset ex_in_conv open_interior subset_eq) qed moreover from \I \ K \ I = I\ have "open ?G" unfolding I_def K_def by (auto simp: fun_eq_iff o_apply) moreover have "?G \ X" proof - have "?G \ K ?G" unfolding K_def using closure_subset by fastforce also from \I \ K \ I = I\ have "... = K {x}" unfolding I_def K_def by (metis closure_interior comp_def double_complement) also from \closed X\ \x \ X\ have "... \ X" unfolding K_def by clarsimp (meson closure_minimal contra_subsetD empty_subsetI insert_subset) finally show ?thesis . qed ultimately show "\T. open T \ x \ T \ T \ X" by blast qed lemma part_I_K_I: assumes "I \ K \ I = (I::'a::topological_space set \ 'a set)" shows "I \ K = (K::'a set \ 'a set)" using interior_open[OF part_closed_open[OF assms closed_closure]] unfolding I_def K_def o_def by simp lemma part_K_I_I: assumes "I \ K \ I = (I::'a::topological_space set \ 'a set)" shows "K \ I = (I::'a set \ 'a set)" using part_I_K_I[OF assms] assms by simp lemma kuratowski_neqs: assumes "kuratowski (X :: 'a::topological_space set)" shows "I \ I \ K \ (I :: 'a set \ 'a set)" "I \ K \ I \ K \ (I :: 'a set \ 'a set)" "I \ K \ I \ I \ (K :: 'a set \ 'a set)" "I \ K \ K \ I \ (K :: 'a set \ 'a set)" "K \ I \ K \ I \ (K :: 'a set \ 'a set)" "K \ I \ K \ (K :: 'a set \ 'a set)" "I \ K \ K \ (I :: 'a set \ 'a set)" "I \ (id :: 'a set \ 'a set)" "K \ (id :: 'a set \ 'a set)" "I \ K \ I \ (id :: 'a set \ 'a set)" "K \ I \ K \ (id :: 'a set \ 'a set)" using assms unfolding kuratowski_def extremally_disconnected_def open_unresolvable_def by (metis (no_types, lifting) I_K K_K I_K_I_K K_I_K_I part_I_K_I part_K_I_I o_assoc comp_id)+ lemma kuratowski_card: assumes "kuratowski (X :: 'a::topological_space set)" shows "card {f. CK_nf_pos (f::'a set \ 'a set)} = 7" using CK_nf_pos_lattice[where 'a='a] kuratowski_neqs[OF assms] assms unfolding kuratowski_def extremally_disconnected_def open_unresolvable_def by (subst CK_nf_pos_set) (subst card_insert_disjoint; (auto)[1])+ text\ \\\ is a Kuratowski space. \ lemma kuratowski_reals: shows "kuratowski (\ :: real set)" unfolding kuratowski_def extremally_disconnected_def open_unresolvable_def by (rule conjI) (metis (no_types, lifting) I_K list.inject nf_RRR(11) nf_RRR(8) o_assoc, metis (no_types, lifting) I_K fun.map_comp list.inject nf_RRR(11) nf_RRR(9)) subsection\ Chagrov's theorem \ theorem chagrov: fixes X :: "'a::topological_space set" obtains "discrete X" | "\discrete X \ part X" | "\discrete X \ ed_ou X" | "\ed_ou X \ open_unresolvable X" | "\ed_ou X \ \part X \ extremally_disconnected X" | "kuratowski X" unfolding kuratowski_def by metis corollary chagrov_card: shows "card {f. CK_nf_pos (f::'a::topological_space set \ 'a set)} \ {1,3,4,5,7}" using discrete_card part_not_discrete_card ed_ou_not_discrete_card open_unresolvable_not_ed_ou_card extremally_disconnected_not_part_not_ed_ou_card kuratowski_card by (cases rule: chagrov) blast+ (*<*) end (*>*) diff --git a/thys/Lower_Semicontinuous/Lower_Semicontinuous.thy b/thys/Lower_Semicontinuous/Lower_Semicontinuous.thy --- a/thys/Lower_Semicontinuous/Lower_Semicontinuous.thy +++ b/thys/Lower_Semicontinuous/Lower_Semicontinuous.thy @@ -1,1859 +1,1859 @@ (* Title: thys/Lower_Semicontinuous/Lower_Semicontinuous.thy Author: Bogdan Grechuk, University of Edinburgh *) section \Lower semicontinuous functions\ theory Lower_Semicontinuous -imports "HOL-Analysis.Ordered_Euclidean_Space" +imports "HOL-Analysis.Multivariate_Analysis" begin subsection\Relative interior in one dimension\ lemma rel_interior_ereal_semiline: fixes a :: ereal shows "rel_interior {y. a \ ereal y} = {y. a < ereal y}" proof (cases a) case (real r) then show ?thesis using rel_interior_real_semiline[of r] by (simp add: atLeast_def greaterThan_def) next case PInf thus ?thesis using rel_interior_empty by auto next case MInf thus ?thesis using rel_interior_UNIV by auto qed lemma closed_ereal_semiline: fixes a :: ereal shows "closed {y. a \ ereal y}" proof (cases a) case (real r) then show ?thesis using closed_real_atLeast unfolding atLeast_def by simp qed auto lemma ereal_semiline_unique: fixes a b :: ereal shows "{y. a \ ereal y} = {y. b \ ereal y} \ a = b" by (metis mem_Collect_eq ereal_le_real order_antisym) subsection \Lower and upper semicontinuity\ definition lsc_at :: "'a \ ('a::topological_space \ 'b::order_topology) \ bool" where "lsc_at x0 f \ (\X l. X \ x0 \ (f \ X) \ l \ f x0 \ l)" definition usc_at :: "'a \ ('a::topological_space \ 'b::order_topology) \ bool" where "usc_at x0 f \ (\X l. X \ x0 \ (f \ X) \ l \ l \ f x0)" lemma lsc_at_mem: assumes "lsc_at x0 f" assumes "x \ x0" assumes "(f \ x) \ A" shows "f x0 \ A" using assms lsc_at_def[of x0 f] by blast lemma usc_at_mem: assumes "usc_at x0 f" assumes "x \ x0" assumes "(f \ x) \ A" shows "f x0 \ A" using assms usc_at_def[of x0 f] by blast lemma lsc_at_open: fixes f :: "'a::first_countable_topology \ 'b::{complete_linorder, linorder_topology}" shows "lsc_at x0 f \ (\S. open S \ f x0 \ S \ (\T. open T \ x0 \ T \ (\x'\T. f x' \ f x0 \ f x' \ S)))" (is "?lhs \ ?rhs") proof- { assume "~?rhs" from this obtain S where S_def: "open S \ f x0 : S \ (\T. (open T \ x0 \ T) \ (\x'\T. f x' \ f x0 \ f x' \ S))" by metis define X where "X = {x'. f x' \ f x0 \ f x' \ S}" hence "x0 islimpt X" unfolding islimpt_def using S_def by auto from this obtain x where x_def: "(\n. x n \ X) \ x \ x0" using islimpt_sequential[of x0 X] by auto hence not: "~(f \ x) \ (f x0)" unfolding lim_explicit using X_def S_def by auto from compact_complete_linorder[of "f \ x"] obtain l r where r_def: "strict_mono r \ ((f \ x) \ r) \ l" by auto { assume "l : S" hence "\N. \n\N. f(x(r n)) \ S" using r_def lim_explicit[of "f \ x \ r" l] S_def by auto hence False using x_def X_def by auto } hence l_prop: "l \ S \ l\f x0" using r_def x_def X_def Lim_bounded[of "f \ x \ r"] by auto { assume "f x0 \ l" hence "f x0 = l" using l_prop by auto hence False using l_prop S_def by auto } hence "\x l. x \ x0 \ (f \ x) \ l \ ~(f x0 \ l)" apply(rule_tac x="x \ r" in exI) apply(rule_tac x=l in exI) using r_def x_def by (auto simp add: o_assoc LIMSEQ_subseq_LIMSEQ) hence "~?lhs" unfolding lsc_at_def by blast } moreover { assume "?rhs" { fix x A assume x_def: "x \ x0" "(f \ x) \ A" { assume "A \ f x0" from this obtain S V where SV_def: "open S \ open V \ f x0 : S \ A : V \ S Int V = {}" using hausdorff[of "f x0" A] by auto from this obtain T where T_def: "open T \ x0 : T \ (\x'\T. (f x' \ f x0 \ f x' \ S))" using \?rhs\ by metis from this obtain N1 where "\n\N1. x n \ T" using x_def lim_explicit[of x x0] by auto hence *: "\n\N1. (f (x n) \ f x0 \ f(x n) \ S)" using T_def by auto from SV_def obtain N2 where "\n\N2. f(x n) \ V" using lim_explicit[of "f \ x" A] x_def by auto hence "\n\(max N1 N2). \(f(x n) \ f x0)" using SV_def * by auto hence "\n\(max N1 N2). f(x n) \ f x0" by auto hence "f x0 \ A" using Lim_bounded2[of "f \ x" A "max N1 N2" "f x0"] x_def by auto } hence "f x0 \ A" by auto } hence "?lhs" unfolding lsc_at_def by blast } ultimately show ?thesis by blast qed lemma lsc_at_open_mem: fixes f :: "'a::first_countable_topology \ 'b::{complete_linorder, linorder_topology}" assumes "lsc_at x0 f" assumes "open S \ f x0 : S" obtains T where "open T \ x0 \ T \ (\x'\T. (f x' \ f x0 \ f x' \ S))" using assms lsc_at_open[of x0 f] by blast lemma lsc_at_MInfty: fixes f :: "'a::topological_space \ ereal" assumes "f x0 = -\" shows "lsc_at x0 f" unfolding lsc_at_def using assms by auto lemma lsc_at_PInfty: fixes f :: "'a::metric_space \ ereal" assumes "f x0 = \" shows "lsc_at x0 f \ continuous (at x0) f" unfolding lsc_at_open continuous_at_open using assms by auto lemma lsc_at_real: fixes f :: "'a::metric_space \ ereal" assumes "\f x0\ \ \" shows "lsc_at x0 f \ (\e. e>0 \ (\T. open T \ x0 : T \ (\y \ T. f y > f x0 - e)))" (is "?lhs \ ?rhs") proof- obtain m where m_def: "f x0 = ereal m" using assms by (cases "f x0") auto { assume lsc: "lsc_at x0 f" { fix e assume e_def: "(e :: ereal)>0" hence *: "f x0 : {f x0 - e <..< f x0 + e}" using assms ereal_between by auto from this obtain T where T_def: "open T \ x0 : T \ (\x'\T. f x' \ f x0 \ f x' \ {f x0 - e <..< f x0 + e})" apply (subst lsc_at_open_mem[of x0 f "{f x0 - e <..< f x0 + e}"]) using lsc e_def by auto { fix y assume "y:T" { assume "f y \ f x0" hence "f y > f x0 - e" using T_def \y:T\ by auto } moreover { assume "f y > f x0" hence "\>f x0 - e" using * by auto } ultimately have "f y > f x0 - e" using not_le by blast } hence "\T. open T \ x0 \ T \ (\y \ T. f y > f x0 - e)" using T_def by auto } hence "?rhs" by auto } moreover { assume "?rhs" { fix S assume S_def: "open S \ f x0 : S" from this obtain e where e_def: "e>0 \ {f x0 - e<.. S" apply (subst ereal_open_cont_interval[of S "f x0"]) using assms by auto from this obtain T where T_def: "open T \ x0 : T \ (\y \ T. f y > f x0 - e)" using \?rhs\[rule_format, of e] by auto { fix y assume "y:T" "f y \ f x0" hence "f y < f x0 + e" using assms e_def ereal_between[of "f x0" e] by auto hence "f y \ S" using e_def T_def \y\T\ by auto } hence "\T. open T \ x0 : T \ (\y\T. (f y \ f x0 \ f y \ S))" using T_def by auto } hence "lsc_at x0 f" using lsc_at_open by auto } ultimately show ?thesis by auto qed lemma lsc_at_ereal: fixes f :: "'a::metric_space \ ereal" shows "lsc_at x0 f \ (\CT. open T \ x0 \ T \ (\y \ T. f y > C))" (is "?lhs \ ?rhs") proof- { assume "f x0 = -\" hence ?thesis using lsc_at_MInfty by auto } moreover { assume pinf: "f x0 = \" { assume lsc: "lsc_at x0 f" { fix C assume "C f x0 : {C<..}" by auto from this obtain T where T_def: "open T \ x0 \ T \ (\y\T. f y \ {C<..})" using pinf lsc lsc_at_PInfty[of f x0] unfolding continuous_at_open by metis hence "\T. open T \ x0 \ T \ (\y\T. C < f y)" by auto } hence "?rhs" by auto } moreover { assume "?rhs" { fix S assume S_def: "open S \ f x0 : S" then obtain C where C_def: "ereal C < f x0 \ {ereal C<..} \ S" using pinf open_PInfty by auto then obtain T where T_def: "open T \ x0 : T \ (\y \ T. f y \ S)" using \?rhs\[rule_format, of "ereal C"] by auto hence "\T. open T \ x0 \ T \ (\y\T. (f y \ f x0 \ f y \ S))" using T_def by auto } hence "lsc_at x0 f" using lsc_at_open by auto } ultimately have ?thesis by blast } moreover { assume fin: "f x0 \ -\" "f x0 \ \" { assume lsc: "lsc_at x0 f" { fix C assume "C0" using fin ereal_less_minus_iff by auto from this obtain T where T_def: "open T \ x0 \ T \ (\y\T. f x0 - (f x0-C) < f y)" using lsc_at_real[of f x0] lsc fin by auto moreover have "f x0 - (f x0-C) = C" using fin apply (cases "f x0", cases C) by auto ultimately have "\T. open T \ x0 \ T \ (\y\T. C < f y)" by auto } hence "?rhs" by auto } moreover { assume "?rhs" { fix e :: ereal assume "e>0" hence "f x0 - e < f x0" using fin apply (cases "f x0", cases e) by auto hence "\T. open T \ x0 \ T \ (\y\T. f x0 - e < f y)" using fin \?rhs\ by auto } hence "lsc_at x0 f" using lsc_at_real[of f x0] fin by auto } ultimately have ?thesis by blast } ultimately show ?thesis by blast qed lemma lst_at_ball: fixes f :: "'a::metric_space => ereal" shows "lsc_at x0 f \ (\Cd>0. \y \ (ball x0 d). C ?rhs") proof- { assume lsc: "lsc_at x0 f" { fix C :: ereal assume "C x0 : T \ (\y \ T. C < f y)" using lsc lsc_at_ereal[of x0 f] by auto hence "\d. d>0 \ (\y \ (ball x0 d). C < f y)" by (force simp add: open_contains_ball) } } moreover { assume "?rhs" { fix C :: ereal assume "C0 \ (\y \ (ball x0 d). C < f y)" using \?rhs\ by auto hence "\T. open T \ x0 \ T \ (\y \ T. C < f y)" apply (rule_tac x="ball x0 d" in exI) apply (simp add: centre_in_ball) done } hence "lsc_at x0 f" using lsc_at_ereal[of x0 f] by auto } ultimately show ?thesis by auto qed lemma lst_at_delta: fixes f :: "'a::metric_space \ ereal" shows "lsc_at x0 f \ (\Cd>0. \y. dist x0 y < d \ C < f y)" (is "?lhs \ ?rhs") proof- have "?rhs \ (\Cd>0. \y \ (ball x0 d). C < f y)" unfolding ball_def by auto thus ?thesis using lst_at_ball[of x0 f] by auto qed lemma lsc_liminf_at: fixes f :: "'a::metric_space \ ereal" shows "lsc_at x0 f \ f x0 \ Liminf (at x0) f" unfolding lst_at_ball le_Liminf_iff eventually_at by (intro arg_cong[where f=All] imp_cong refl ext ex_cong) (auto simp: dist_commute zero_less_dist_iff) lemma lsc_liminf_at_eq: fixes f :: "'a::metric_space \ ereal" shows "lsc_at x0 f \ (f x0 = min (f x0) (Liminf (at x0) f))" by (metis inf_ereal_def le_iff_inf lsc_liminf_at) lemma lsc_imp_liminf: fixes f :: "'a::metric_space \ ereal" assumes "lsc_at x0 f" assumes "x \ x0" shows "f x0 \ liminf (f \ x)" proof (cases "f x0") case PInf then show ?thesis using assms lsc_at_PInfty[of f x0] lim_imp_Liminf[of _ "f \ x"] continuous_at_sequentially[of x0 f] by auto next case (real r) { fix e assume e_def: "(e :: ereal)>0" from this obtain T where T_def: "open T \ x0 : T \ (\y \ T. f y > f x0 - e)" using lsc_at_real[of f x0] real assms by auto from this obtain N where N_def: "\n\N. x n \ T" apply (subst tendsto_obtains_N[of x x0 T]) using assms by auto hence "\n\N. f x0 - e < (f \ x) n" using T_def by auto hence "liminf (f \ x) \ f x0 - e" by (intro Liminf_bounded) (auto simp: eventually_sequentially intro!: exI[of _ N]) hence "f x0 \ liminf (f \ x) + e" apply (cases e) unfolding ereal_minus_le_iff by auto } then show ?thesis by (intro ereal_le_epsilon) auto qed auto lemma lsc_liminf: fixes f :: "'a::metric_space \ ereal" shows "lsc_at x0 f \ (\x. x \ x0 \ f x0 \ liminf (f \ x))" (is "?lhs \ ?rhs") proof- { assume "?rhs" { fix x A assume x_def: "x \ x0" "(f \ x) \ A" hence "f x0 \ A" using \?rhs\ lim_imp_Liminf[of sequentially] by auto } hence "?lhs" unfolding lsc_at_def by blast } from this show ?thesis using lsc_imp_liminf by auto qed lemma lsc_sequentially: fixes f :: "'a::metric_space \ ereal" shows "lsc_at x0 f \ (\x c. x \ x0 \ (\n. f(x n)\c) \ f(x0)\c)" (is "?lhs \ ?rhs") proof- { assume "?rhs" { fix x l assume "x \ x0" "(f \ x) \ l" { assume "l = \" hence "f x0 \ l" by auto } moreover { assume "l = -\" { fix B :: real obtain N where N_def: "\n\N. f(x n) \ ereal B" using Lim_MInfty[of "f \ x"] \(f \ x) \ l\ \l = -\\ by auto define g where "g n = (if n\N then x n else x N)" for n hence "g \ x0" by (intro filterlim_cong[THEN iffD1, OF refl refl _ \x \ x0\]) (auto simp: eventually_sequentially) moreover have "\n. f(g n) \ ereal B" using g_def N_def by auto ultimately have "f x0 \ ereal B" using \?rhs\ by auto } hence "f x0 = -\" using ereal_bot by auto hence "f x0 \ l" by auto } moreover { assume fin: "\l\ \ \" { fix e assume e_def: "(e :: ereal)>0" from this obtain N where N_def: "\n\N. f(x n) \ {l - e <..< l + e}" apply (subst tendsto_obtains_N[of "f \ x" l "{l - e <..< l + e}"]) using fin e_def ereal_between \(f \ x) \ l\ by auto define g where "g n = (if n\N then x n else x N)" for n hence "g \ x0" by (intro filterlim_cong[THEN iffD1, OF refl refl _ \x \ x0\]) (auto simp: eventually_sequentially) moreover have "\n. f(g n) \ l + e" using g_def N_def by auto ultimately have "f x0 \ l + e" using \?rhs\ by auto } hence "f x0 \ l" using ereal_le_epsilon by auto } ultimately have "f x0 \ l" by blast } hence "lsc_at x0 f" unfolding lsc_at_def by auto } moreover { assume lsc: "lsc_at x0 f" { fix x c assume xc_def: "x \ x0 \ (\n. f(x n)\c)" hence "liminf (f \ x) \ c" using Limsup_bounded[where F = sequentially and X = "f \ x" and C = c] Liminf_le_Limsup[of sequentially "f \ x"] by auto hence "f x0 \ c" using lsc xc_def lsc_imp_liminf[of x0 f x] by auto } hence "?rhs" by auto } ultimately show ?thesis by blast qed lemma lsc_sequentially_gen: fixes f :: "'a::metric_space \ ereal" shows "lsc_at x0 f \ (\x c c0. x \ x0 \ c \ c0 \ (\n. f(x n)\c n) \ f(x0)\c0)" (is "?lhs \ ?rhs") proof- { assume "?rhs" { fix x c0 assume a: "x \ x0 \ (\n. f (x n) \ c0)" define c where "c = (\n::nat. c0)" hence "c \ c0" by auto hence "f(x0)\c0" using \?rhs\[rule_format, of x c c0] using a c_def by auto } hence "?lhs" using lsc_sequentially[of x0 f] by auto } moreover { assume lsc: "lsc_at x0 f" { fix x c c0 assume xc_def: "x \ x0 \ c \ c0 \ (\n. f(x n)\c n)" hence "liminf (f \ x) \ c0" using Liminf_mono[of "f \ x" c sequentially] lim_imp_Liminf[of sequentially] by auto hence "f x0 \ c0" using lsc xc_def lsc_imp_liminf[of x0 f x] by auto } hence "?rhs" by auto } ultimately show ?thesis by blast qed lemma lsc_sequentially_mem: fixes f :: "'a::metric_space \ ereal" assumes "lsc_at x0 f" assumes "x \ x0" "c \ c0" assumes "\n. f(x n)\c n" shows "f(x0)\c0" using lsc_sequentially_gen[of x0 f] assms by auto lemma lsc_uminus: fixes f :: "'a::metric_space \ ereal" shows "lsc_at x0 (\x. -f x) \ usc_at x0 f" proof- { assume lsc: "lsc_at x0 (\x. -f x)" { fix x A assume x_def: "x \ x0" "(f \ x) \ A" hence "(\i. - f (x i)) \ -A" using tendsto_uminus_ereal[of "f \ x" A] by auto hence "((\x. - f x) \ x) \ -A" unfolding o_def by auto hence " - f x0 \ - A" apply (subst lsc_at_mem[of x0 "(\x. -f x)" x]) using lsc x_def by auto hence "f x0 \ A" by auto } hence "usc_at x0 f" unfolding usc_at_def by auto } moreover { assume usc: "usc_at x0 f" { fix x A assume x_def: "x \ x0" "((\x. - f x) \ x) \ A" hence "(\i. - f (x i)) \ A" unfolding o_def by auto hence "(\i. f (x i)) \ - A" using tendsto_uminus_ereal[of "(\i. - f (x i))" A] by auto hence "(f \ x) \ -A" unfolding o_def by auto hence "f x0 \ - A" apply (subst usc_at_mem[of x0 "f" x]) using usc x_def by auto hence "-f x0 \ A" by (auto simp: ereal_uminus_le_reorder) } hence "lsc_at x0 (\x. -f x)" unfolding lsc_at_def by auto } ultimately show ?thesis by blast qed lemma usc_limsup: fixes f :: "'a::metric_space \ ereal" shows "usc_at x0 f \ (\x. x \ x0 \ f x0 \ limsup (f \ x))" (is "?lhs \ ?rhs") proof- have "usc_at x0 f \ (\x. x \ x0 \ - f x0 \ liminf ((\x. - f x) \ x))" using lsc_uminus[of x0 f] lsc_liminf[of x0 "(\x. - f x)"] by auto moreover { fix x assume "x \ x0" have "(-f x0 \ -limsup (f \ x)) \ (-f x0 \ liminf ((\x. - f x) \ x))" using ereal_Liminf_uminus[of _ "f \ x"] unfolding o_def by auto hence "(f x0 \ limsup (f \ x)) \ (-f x0 \ liminf ((\x. - f x) \ x))" by auto } ultimately show ?thesis by auto qed lemma usc_imp_limsup: fixes f :: "'a::metric_space \ ereal" assumes "usc_at x0 f" assumes "x \ x0" shows "f x0 \ limsup (f \ x)" using assms usc_limsup[of x0 f] by auto lemma usc_limsup_at: fixes f :: "'a::metric_space \ ereal" shows "usc_at x0 f \ f x0 \ Limsup (at x0) f" proof- have "usc_at x0 f \ lsc_at x0 (\x. -(f x))" by (metis lsc_uminus) also have "\ \ -(f x0) \ Liminf (at x0) (\x. -(f x))" by (metis lsc_liminf_at) also have "\ \ -(f x0) \ -(Limsup (at x0) f)" by (metis ereal_Liminf_uminus) finally show ?thesis by auto qed lemma continuous_iff_lsc_usc: fixes f :: "'a::metric_space \ ereal" shows "continuous (at x0) f \ (lsc_at x0 f) \ (usc_at x0 f)" proof- { assume a: "continuous (at x0) f" { fix x assume "x \ x0" hence "(f \ x) \ f x0" using a continuous_imp_tendsto[of x0 f x] by auto hence "liminf (f \ x) = f x0 \ limsup (f \ x) = f x0" using lim_imp_Liminf[of sequentially] lim_imp_Limsup[of sequentially] by auto } hence "lsc_at x0 f \ usc_at x0 f" unfolding lsc_liminf usc_limsup by auto } moreover { assume a: "(lsc_at x0 f) \ (usc_at x0 f)" { fix x assume "x \ x0" hence "limsup (f \ x) \ f x0" using a unfolding usc_limsup by auto moreover have "\ \ liminf (f \ x)" using a \x \ x0\ unfolding lsc_liminf by auto ultimately have "limsup (f \ x) = f x0 \ liminf (f \ x) = f x0" using Liminf_le_Limsup[of sequentially "f \ x"] by auto hence "(f \ x) \ f x0" using Liminf_eq_Limsup[of sequentially] by auto } hence "continuous (at x0) f" using continuous_at_sequentially[of x0 f] by auto } ultimately show ?thesis by blast qed lemma continuous_lsc_compose: assumes "lsc_at (g x0) f" "continuous (at x0) g" shows "lsc_at x0 (f \ g)" proof- { fix x L assume "x \ x0" "(f \ g \ x) \ L" hence "f(g x0) \ L" apply (subst lsc_at_mem[of "g x0" f "g \ x" L]) using assms continuous_imp_tendsto[of x0 g x] unfolding o_def by auto } from this show ?thesis unfolding lsc_at_def by auto qed lemma continuous_isCont: "continuous (at x0) f \ isCont f x0" by (metis continuous_at isCont_def) lemma isCont_iff_lsc_usc: fixes f :: "'a::metric_space \ ereal" shows "isCont f x0 \ (lsc_at x0 f) \ (usc_at x0 f)" by (metis continuous_iff_lsc_usc continuous_isCont) definition lsc :: "('a::topological_space \ 'b::order_topology) \ bool" where "lsc f \ (\x. lsc_at x f)" definition usc :: "('a::topological_space \ 'b::order_topology) \ bool" where "usc f \ (\x. usc_at x f)" lemma continuous_UNIV_iff_lsc_usc: fixes f :: "'a::metric_space \ ereal" shows "(\x. continuous (at x) f) \ (lsc f) \ (usc f)" by (metis continuous_iff_lsc_usc lsc_def usc_def) subsection \Epigraphs\ definition "Epigraph S (f::_ \ ereal) = {xy. fst xy : S \ f (fst xy) \ ereal(snd xy)}" lemma mem_Epigraph: "(x, y) \ Epigraph S f \ x \ S \ f x \ ereal y" unfolding Epigraph_def by auto lemma ereal_closed_levels: fixes f :: "'a::metric_space \ ereal" shows "(\y. closed {x. f(x)\y}) \ (\r. closed {x. f(x)\ereal r})" (is "?lhs \ ?rhs") proof- { assume "?rhs" { fix y :: ereal { assume "y \ \ \ y \ -\" hence "closed {x. f(x)\y}" using \?rhs\ by (cases y) auto } moreover { assume "y=\" hence "closed {x. f(x)\y}" by auto } moreover { assume "y=-\" hence "{x. f(x)\y} = Inter {{x. f(x)\ereal r} |r. r : UNIV}" using ereal_bot by auto hence "closed {x. f(x)\y}" using closed_Inter \?rhs\ by auto } ultimately have "closed {x. f(x)\y}" by blast } hence "?lhs" by auto } from this show ?thesis by auto qed lemma lsc_iff: fixes f :: "'a::metric_space \ ereal" shows "(lsc f \ (\y. closed {x. f(x)\y})) \ (lsc f \ closed (Epigraph UNIV f))" proof- { assume "lsc f" { fix z z0 assume a: "\n. z n \ (Epigraph UNIV f) \ z \ z0" { fix n have "z n : (Epigraph UNIV f)" using a by auto hence "f (fst (z n)) \ ereal(snd (z n))" using a unfolding Epigraph_def by auto hence "\xn cn. z n = (xn, cn) \ f(xn)\ereal cn" apply (rule_tac x="fst (z n)" in exI) apply (rule_tac x="snd (z n)" in exI) by auto } from this obtain x c where xc_def: "\n. z n = (x n, c n) \ f(x n)\ereal (c n)" by metis from this a have "\x0 c0. z0 = (x0, c0) \ x \ x0 \ c \ c0" apply (rule_tac x="fst z0" in exI) apply (rule_tac x="snd z0" in exI) using tendsto_fst[of z z0] tendsto_snd[of z z0] by auto from this obtain x0 c0 where xc0_def: "z0 = (x0, c0) \ x \ x0 \ c \ c0" by auto hence "f(x0)\ereal c0" apply (subst lsc_sequentially_mem[of x0 f x "ereal \ c" "ereal c0"]) using \lsc f\ xc_def unfolding lsc_def unfolding o_def by auto hence "z0 : (Epigraph UNIV f)" unfolding Epigraph_def using xc0_def by auto } hence "closed (Epigraph UNIV f)" by (simp add: closed_sequential_limits) } moreover { assume "closed (Epigraph UNIV f)" hence *: "\x l. (\n. f (fst (x n)) \ ereal(snd (x n))) \ x \ l \ f (fst l) \ ereal(snd l)" unfolding Epigraph_def closed_sequential_limits by auto { fix r :: real { fix z z0 assume a: "\n. f(z n)\ereal r \ z \ z0" hence "f(z0)\ereal r" using *[rule_format, of "(\n. (z n,r))" "(z0,r)"] tendsto_Pair[of z z0] by auto } hence "closed {x. f(x)\ereal r}" by (simp add: closed_sequential_limits) } hence "\y. closed {x. f(x)\y}" using ereal_closed_levels by auto } moreover { assume a: "\y. closed {x. f(x)\ y}" { fix x0 { fix x l assume "x \ x0" "(f \ x) \ l" { assume "l = \" hence "f x0 \ l" by auto } moreover { assume mi: "l = -\" { fix B :: real obtain N where N_def: "\n\N. f(x n)\ereal B" using mi \(f \ x) \ l\ Lim_MInfty[of "f \ x"] by auto { fix d assume "(d::real)>0" from this obtain N1 where N1_def: "\n\N1. dist (x n) x0 < d" using \x \ x0\ unfolding lim_sequentially by auto hence "\y. dist y x0 < d \ y : {x. f(x)\ereal B}" apply (rule_tac x="x (max N N1)" in exI) using N_def by auto } hence "x0 : closure {x. f(x)\ereal B}" apply (subst closure_approachable) by auto hence "f x0 \ ereal B" using a by auto } hence "f x0 \ l" using ereal_bot[of "f x0"] by auto } moreover { assume fin: "\l\ \ \" { fix e assume e_def: "(e :: ereal)>0" from this obtain N where N_def: "\n\N. f(x n) : {l - e <..< l + e}" apply (subst tendsto_obtains_N[of "f \ x" l "{l - e <..< l + e}"]) using fin e_def ereal_between \(f \ x) \ l\ by auto hence *: "\n\N. x n : {x. f(x)\l+e}" using N_def by auto { fix d assume "(d::real)>0" from this obtain N1 where N1_def: "\n\N1. dist (x n) x0 < d" using \x \ x0\ unfolding lim_sequentially by auto hence "\y. dist y x0 < d \ y : {x. f(x)\l+e}" apply (rule_tac x="x (max N N1)" in exI) using * by auto } hence "x0 : closure {x. f(x)\l+e}" apply (subst closure_approachable) by auto hence "f x0 \ l+e" using a by auto } hence "f x0 \ l" using ereal_le_epsilon by auto } ultimately have "f x0 \ l" by blast } hence "lsc_at x0 f" unfolding lsc_at_def by auto } hence "lsc f" unfolding lsc_def by auto } ultimately show ?thesis by auto qed definition lsc_hull :: "('a::metric_space \ ereal) \ ('a::metric_space \ ereal)" where "lsc_hull f = (SOME g. Epigraph UNIV g = closure(Epigraph UNIV f))" lemma epigraph_mono: fixes f :: "'a::metric_space \ ereal" shows "(x,y):Epigraph UNIV f \ y\z \ (x,z):Epigraph UNIV f" unfolding Epigraph_def apply auto using ereal_less_eq(3)[of y z] order_trans_rules(23) by blast lemma closed_epigraph_lines: fixes S :: "('a::metric_space * 'b::metric_space) set" assumes "closed S" shows "closed {z. (x, z) : S}" proof- { fix z assume z_def: "z : closure {z. (x, z) : S}" { fix e :: real assume "e>0" from this obtain y where y_def: "(x,y) : S \ dist y z < e" using closure_approachable[of z "{z. (x, z) : S}"] z_def by auto moreover have "dist (x,y) (x,z) = dist y z" unfolding dist_prod_def by auto ultimately have "\s. s \ S \ dist s (x,z) < e" apply(rule_tac x="(x,y)" in exI) by auto } hence "(x,z):S" using closed_approachable[of S "(x,z)"] assms by auto } hence "closure {z. (x, z) : S} \ {z. (x, z) : S}" by blast from this show ?thesis using closure_subset_eq by auto qed lemma mono_epigraph: fixes S :: "('a::metric_space * real) set" assumes mono: "\x y z. (x,y):S \ y\z \ (x,z):S" assumes "closed S" shows "\g. ((Epigraph UNIV g) = S)" proof- { fix x have "closed {z. (x, z) : S}" using \closed S\ closed_epigraph_lines by auto hence "\a. {z. (x, z) : S} = {z. a \ ereal z}" apply (subst mono_closed_ereal) using mono by auto } from this obtain g where g_def: "\x. {z. (x, z) : S} = {z. g x \ ereal z}" by metis { fix s have "s:S \ (fst s, snd s) : S" by auto also have "\ \ g(fst s) \ ereal (snd s)" using g_def[rule_format, of "fst s"] by blast finally have "s:S \ g(fst s) \ ereal (snd s)" by auto } hence "(Epigraph UNIV g) = S" unfolding Epigraph_def by auto from this show ?thesis by auto qed lemma lsc_hull_exists: fixes f :: "'a::metric_space \ ereal" shows "\g. Epigraph UNIV g = closure (Epigraph UNIV f)" proof- { fix x y z assume xy: "(x,y):closure (Epigraph UNIV f) \ y\z" { fix e :: real assume "e>0" hence "\ya\Epigraph UNIV f. dist ya (x, y) < e" using xy closure_approachable[of "(x,y)" "Epigraph UNIV f"] by auto from this obtain a b where ab: "(a,b):Epigraph UNIV f \ dist (a,b) (x,y) < e" by auto moreover have "dist (a,b) (x,y) = dist (a,b+(z-y)) (x,z)" unfolding dist_prod_def dist_norm by (simp add: algebra_simps) moreover have "(a,b+(z-y)):Epigraph UNIV f" apply (subst epigraph_mono[of _ b]) using ab xy by auto ultimately have "\w\Epigraph UNIV f. dist w (x, z) < e" by auto } hence "(x,z):closure (Epigraph UNIV f)" using closure_approachable by auto } hence "\x y z. (x,y) \ closure (Epigraph UNIV f) \ y\z \ (x,z) \ closure (Epigraph UNIV f)" by auto from this show ?thesis using mono_epigraph[of "closure (Epigraph UNIV f)"] by auto qed lemma epigraph_invertible: assumes "Epigraph UNIV f = Epigraph UNIV g" shows "f=g" proof- { fix x { fix C have "f x \ ereal C \ (x,C) : Epigraph UNIV f" unfolding Epigraph_def by auto also have "\ \ (x,C) : Epigraph UNIV g" using assms by auto also have "\ \ g x \ ereal C" unfolding Epigraph_def by auto finally have "f x \ ereal C \ g x \ ereal C" by auto } hence "g x = f x" using ereal_le_real[of "g x" "f x"] ereal_le_real[of "f x" "g x"] by auto } from this show ?thesis by (simp add: ext) qed lemma lsc_hull_ex_unique: fixes f :: "'a::metric_space \ ereal" shows "\!g. Epigraph UNIV g = closure (Epigraph UNIV f)" using lsc_hull_exists epigraph_invertible by metis lemma epigraph_lsc_hull: fixes f :: "'a::metric_space \ ereal" shows "Epigraph UNIV (lsc_hull f) = closure(Epigraph UNIV f)" proof- have "\g. Epigraph UNIV g = closure (Epigraph UNIV f)" using lsc_hull_exists by auto thus ?thesis unfolding lsc_hull_def using some_eq_ex[of "(\g. Epigraph UNIV g = closure(Epigraph UNIV f))"] by auto qed lemma lsc_hull_expl: "(g = lsc_hull f) \ (Epigraph UNIV g = closure(Epigraph UNIV f))" proof- { assume "Epigraph UNIV g = closure(Epigraph UNIV f)" hence "lsc_hull f = g" unfolding lsc_hull_def apply (subst some1_equality[of _ g]) using lsc_hull_ex_unique by metis+ } from this show ?thesis using epigraph_lsc_hull by auto qed lemma lsc_lsc_hull: "lsc (lsc_hull f)" using epigraph_lsc_hull[of f] lsc_iff[of "lsc_hull f"] by auto lemma epigraph_subset_iff: fixes f g :: "'a::metric_space \ ereal" shows "Epigraph UNIV f \ Epigraph UNIV g \ (\x. g x \ f x)" proof- { assume epi: "Epigraph UNIV f \ Epigraph UNIV g" { fix x { fix z assume "f x \ ereal z" hence "(x,z)\Epigraph UNIV f" unfolding Epigraph_def by auto hence "(x,z)\Epigraph UNIV g" using epi by auto hence "g x \ ereal z" unfolding Epigraph_def by auto } hence "g x \ f x" apply (subst ereal_le_real) by auto } } moreover { assume le: "\x. g x \ f x" { fix x y assume "(x,y):Epigraph UNIV f" hence "f x \ ereal y" unfolding Epigraph_def by auto moreover have "g x \ f x" using le by auto ultimately have "g x \ ereal y" by auto hence "(x,y):Epigraph UNIV g" unfolding Epigraph_def by auto } } ultimately show ?thesis by auto qed lemma lsc_hull_le: "(lsc_hull f) x \ f x" using epigraph_lsc_hull[of f] closure_subset epigraph_subset_iff[of f "lsc_hull f"] by auto lemma lsc_hull_greatest: fixes f g :: "'a::metric_space \ ereal" assumes "lsc g" "\x. g x \ f x" shows "\x. g x \ (lsc_hull f) x" proof- have "closure(Epigraph UNIV f) \ Epigraph UNIV g" using lsc_iff epigraph_subset_iff assms by (metis closure_minimal) from this show ?thesis using epigraph_subset_iff lsc_hull_expl by metis qed lemma lsc_hull_iff_greatest: fixes f g :: "'a::metric_space \ ereal" shows "(g = lsc_hull f) \ lsc g \ (\x. g x \ f x) \ (\h. lsc h \ (\x. h x \ f x) \ (\x. h x \ g x))" (is "?lhs \ ?rhs") proof- { assume "?lhs" hence "?rhs" using lsc_lsc_hull lsc_hull_le lsc_hull_greatest by metis } moreover { assume "?rhs" { fix x have "(lsc_hull f) x \ g x" using \?rhs\ lsc_lsc_hull lsc_hull_le by metis moreover have "(lsc_hull f) x \ g x" using \?rhs\ lsc_hull_greatest by metis ultimately have "(lsc_hull f) x = g x" by auto } hence "?lhs" by (simp add: ext) } ultimately show ?thesis by blast qed lemma lsc_hull_mono: fixes f g :: "'a::metric_space \ ereal" assumes "\x. g x \ f x" shows "\x. (lsc_hull g) x \ (lsc_hull f) x" proof- { fix x have "(lsc_hull g) x \ g x" using lsc_hull_le[of g x] by auto also have "\ \ f x" using assms by auto finally have "(lsc_hull g) x \ f x" by auto } note * = this show ?thesis apply (subst lsc_hull_greatest) using lsc_lsc_hull[of g] * by auto qed lemma lsc_hull_lsc: "lsc f \ (f = lsc_hull f)" using lsc_hull_iff_greatest[of f f] by auto lemma lsc_hull_liminf_at: fixes f :: "'a::metric_space \ ereal" shows "\x. (lsc_hull f) x = min (f x) (Liminf (at x) f)" proof- { fix x z assume "(x,z):Epigraph UNIV (\x. min (f x) (Liminf (at x) f))" hence xz_def: "ereal z \ (SUP e\{0<..}. INF y\ball x e. f y)" unfolding Epigraph_def min_Liminf_at by auto { fix e::real assume "e>0" hence "e/sqrt 2>0" using \e>0\ by simp from this obtain e1 where e1_def: "e1 e1>0" using dense by auto hence "(SUP e\{0<..}. INF y\ball x e. f y) \ (INF y\ball x e1. f y)" by (auto intro: SUP_upper) hence "ereal z \ (INF y\ball x e1. f y)" using xz_def by auto hence *: "\y>ereal z. \t. t \ ball x e1 \ f t \ y" by (simp add: Bex_def Inf_le_iff_less) obtain t where t_def: "t : ball x e1 \ f t \ ereal(z+e1)" using e1_def *[rule_format, of "ereal(z+e1)"] by auto hence epi: "(t,z+e1):Epigraph UNIV f" unfolding Epigraph_def by auto have "dist x t < e1" using t_def unfolding ball_def dist_norm by auto hence "sqrt (e1 ^ 2 + dist x t ^ 2) < e" using e1_def apply (subst real_sqrt_sum_squares_less) by auto moreover have "dist (x,z) (t,z+e1) = sqrt (e1 ^ 2 + dist x t ^ 2)" unfolding dist_prod_def dist_norm by (simp add: algebra_simps) ultimately have "dist (x,z) (t,z+e1) < e" by auto hence "\y\Epigraph UNIV f. dist y (x, z) < e" apply (rule_tac x="(t,z+e1)" in bexI) apply (simp add: dist_commute) using epi by auto } hence "(x,z) : closure (Epigraph UNIV f)" using closure_approachable[of "(x,z)" "Epigraph UNIV f"] by auto } moreover { fix x z assume xz_def: "(x,z):closure (Epigraph UNIV f)" { fix e::real assume "e>0" from this obtain y where y_def: "y:(Epigraph UNIV f) \ dist y (x,z) < e" using closure_approachable[of "(x,z)" "Epigraph UNIV f"] xz_def by metis have "dist (fst y) x \ sqrt ((dist (fst y) x) ^ 2 + (dist (snd y) z) ^ 2)" by (auto intro: real_sqrt_sum_squares_ge1) also have "\ sqrt ((dist (fst y) x) ^ 2 + (dist (snd y) z) ^ 2)" by (auto intro: real_sqrt_sum_squares_ge2) also have "\ball x e. f y) \ f(fst y)" using h1 by (simp add: INF_lower) also have "\\ereal(snd y)" using y_def unfolding Epigraph_def by auto also have "\ < ereal(z+e)" using h2 unfolding dist_norm by auto finally have "(INF y\ball x e. f y) < ereal(z+e)" by auto } hence *: "\e>0. (INF y\ball x e. f y) < ereal(z+e)" by auto { fix e assume "(e::real)>0" { fix d assume "(d::real)>0" { assume "dball x d. f y) < ereal(z+d)" using * \d>0\ by auto also have "\< ereal(z+e)" using \d by auto finally have "(INF y\ball x d. f y) < ereal(z+e)" by auto } moreover { assume "~(d ball x d" by auto hence "(INF y\ball x d. f y) \ (INF y\ball x e. f y)" apply (subst INF_mono) by auto also have "\< ereal(z+e)" using * \e>0\ by auto finally have "(INF y\ball x d. f y) < ereal(z+e)" by auto } ultimately have "(INF y\ball x d. f y) < ereal(z+e)" by blast hence "(INF y\ball x d. f y) \ ereal(z+e)" by auto } hence "min (f x) (Liminf (at x) f) \ ereal (z+e)" unfolding min_Liminf_at by (auto intro: SUP_least) } hence "min (f x) (Liminf (at x) f) \ ereal z" using ereal_le_epsilon2 by auto hence "(x,z):Epigraph UNIV (\x. min (f x) (Liminf (at x) f))" unfolding Epigraph_def by auto } ultimately have "Epigraph UNIV (\x. min (f x) (Liminf (at x) f))= closure (Epigraph UNIV f)" by auto hence "(\x. min (f x) (Liminf (at x) f)) = lsc_hull f" using epigraph_invertible epigraph_lsc_hull[of f] by blast from this show ?thesis by metis qed lemma lsc_hull_same_inf: fixes f :: "'a::metric_space \ ereal" shows "(INF x. lsc_hull f x) = (INF x. f x)" proof- { fix x have "(INF x. f x) \ (INF y\ball x 1. f y)" apply (subst INF_mono) by auto also have "\ \ min (f x) (Liminf (at x) f)" unfolding min_Liminf_at by (auto intro: SUP_upper) also have "\=(lsc_hull f) x" using lsc_hull_liminf_at[of f] by auto finally have "(INF x. f x) \ (lsc_hull f) x" by auto } hence "(INF x. f x) \ (INF x. lsc_hull f x)" apply (subst INF_greatest) by auto moreover have "(INF x. lsc_hull f x) \ (INF x. f x)" apply (subst INF_mono) using lsc_hull_le by auto ultimately show ?thesis by auto qed subsection \Convex Functions\ definition convex_on :: "'a::real_vector set \ ('a \ ereal) \ bool" where "convex_on s f \ (\x\s. \y\s. \u\0. \v\0. u + v = 1 \ f (u *\<^sub>R x + v *\<^sub>R y) \ ereal u * f x + ereal v * f y)" lemma convex_on_ereal_mem: assumes "convex_on s f" assumes "x:s" "y:s" assumes "u\0" "v\0" "u+v=1" shows "f (u *\<^sub>R x + v *\<^sub>R y) \ ereal u * f x + ereal v * f y" using assms unfolding convex_on_def by auto lemma convex_on_ereal_subset: "convex_on t f \ s \ t \ convex_on s f" unfolding convex_on_def by auto lemma convex_on_ereal_univ: "convex_on UNIV f \ (\S. convex_on S f)" using convex_on_ereal_subset by auto lemma ereal_pos_sum_distrib_left: fixes f :: "'a \ ereal" assumes "r\0" "r \ \" shows "r * sum f A = sum (\n. r * f n) A" proof (cases "finite A") case True thus ?thesis proof induct case empty thus ?case by simp next case (insert x A) thus ?case using assms by (simp add: ereal_pos_distrib) qed next case False thus ?thesis by simp qed lemma convex_ereal_add: fixes f g :: "'a::real_vector \ ereal" assumes "convex_on s f" "convex_on s g" shows "convex_on s (\x. f x + g x)" proof- { fix x y assume "x:s" "y:s" moreover fix u v ::real assume uv: "0 \ u" "0 \ v" "u + v = 1" ultimately have "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \ (ereal u * f x + ereal v * f y) + (ereal u * g x + ereal v * g y)" using assms unfolding convex_on_def by (auto simp add: add_mono) also have "\=(ereal u * f x + ereal u * g x) + (ereal v * f y + ereal v * g y)" by (simp add: algebra_simps) also have "\= ereal u * (f x + g x) + ereal v * (f y + g y)" using uv by (simp add: ereal_pos_distrib) finally have "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \ ereal u * (f x + g x) + ereal v * (f y + g y)" by auto } thus ?thesis unfolding convex_on_def by auto qed lemma convex_ereal_cmul: assumes "0 \ (c::ereal)" "convex_on s f" shows "convex_on s (\x. c * f x)" proof- { fix x y assume "x:s" "y:s" moreover fix u v ::real assume uv: "0 \ u" "0 \ v" "u + v = 1" ultimately have "f (u *\<^sub>R x + v *\<^sub>R y) \ (ereal u * f x + ereal v * f y)" using assms unfolding convex_on_def by auto hence "c * f (u *\<^sub>R x + v *\<^sub>R y) \ c * (ereal u * f x + ereal v * f y)" using assms by (intro ereal_mult_left_mono) auto also have "\ \ c * (ereal u * f x) + c * (ereal v * f y)" using assms by (simp add: ereal_le_distrib) also have "\ = ereal u *(c* f x) + ereal v *(c* f y)" by (simp add: algebra_simps) finally have "c * f (u *\<^sub>R x + v *\<^sub>R y) \ ereal u * (c * f x) + ereal v * (c * f y)" by auto } thus ?thesis unfolding convex_on_def by auto qed lemma convex_ereal_max: fixes f g :: "'a::real_vector \ ereal" assumes "convex_on s f" "convex_on s g" shows "convex_on s (\x. max (f x) (g x))" proof- { fix x y assume "x:s" "y:s" moreover fix u v ::real assume uv: "0 \ u" "0 \ v" "u + v = 1" ultimately have "max (f (u *\<^sub>R x + v *\<^sub>R y)) (g (u *\<^sub>R x + v *\<^sub>R y)) \ max (ereal u * f x + ereal v * f y) (ereal u * g x + ereal v * g y)" apply (subst max.mono) using assms unfolding convex_on_def by auto also have "\\ereal u * max (f x) (g x) + ereal v * max (f y) (g y)" apply (subst max.boundedI) apply (subst add_mono) prefer 4 apply (subst add_mono) by (subst ereal_mult_left_mono, auto simp add: uv)+ finally have "max (f (u *\<^sub>R x + v *\<^sub>R y)) (g (u *\<^sub>R x + v *\<^sub>R y)) \ ereal u * max (f x) (g x) + ereal v * max (f y) (g y)" by auto } thus ?thesis unfolding convex_on_def by auto qed lemma convex_on_ereal_alt: fixes C :: "'a::real_vector set" assumes "convex C" shows "convex_on C f = (\x \ C. \y \ C. \m :: real. m \ 0 \ m \ 1 \ f (m *\<^sub>R x + (1 - m) *\<^sub>R y) \ (ereal m) * f x + (1 - (ereal m)) * f y)" proof safe fix x y fix m :: real have[simp]: "ereal (1 - m) = (1 - ereal m)" using ereal_minus(1)[of 1 m] by (auto simp: one_ereal_def) assume asms: "convex_on C f" "x : C" "y : C" "0 \ m" "m \ 1" from this[unfolded convex_on_def, rule_format] have "\u v. ((0 \ u \ 0 \ v \ u + v = 1) \ f (u *\<^sub>R x + v *\<^sub>R y) \ (ereal u) * f x + (ereal v) * f y)" by auto from this[rule_format, of "m" "1 - m", simplified] asms show "f (m *\<^sub>R x + (1 - m) *\<^sub>R y) \ (ereal m) * f x + (1 - ereal m) * f y" by auto next assume asm: "\x\C. \y\C. \m. 0 \ m \ m \ 1 \ f (m *\<^sub>R x + (1 - m) *\<^sub>R y) \ (ereal m) * f x + (1 - ereal m) * f y" { fix x y fix u v :: real assume lasm: "x \ C" "y \ C" "u \ 0" "v \ 0" "u + v = 1" hence[simp]: "1-u=v" "1 - ereal u = ereal v" using ereal_minus(1)[of 1 m] by (auto simp: one_ereal_def) from asm[rule_format, of x y u] have "f (u *\<^sub>R x + v *\<^sub>R y) \ (ereal u) * f x + (ereal v) * f y" using lasm by auto } thus "convex_on C f" unfolding convex_on_def by auto qed lemma convex_on_ereal_alt_mem: fixes C :: "'a::real_vector set" assumes "convex C" assumes "convex_on C f" assumes "x : C" "y : C" assumes "(m::real) \ 0" "m \ 1" shows "f (m *\<^sub>R x + (1 - m) *\<^sub>R y) \ (ereal m) * f x + (1 - (ereal m)) * f y" by (metis assms convex_on_ereal_alt) lemma ereal_add_right_mono: "(a::ereal) \ b \ a + c \ b + c" by (metis add_mono order_refl) lemma convex_on_ereal_sum_aux: assumes "1-a>0" shows "(1 - ereal a) * (ereal (c / (1 - a)) * b) = (ereal c) * b" by (metis mult.assoc mult.commute eq_divide_eq ereal_minus(1) assms one_ereal_def less_le times_ereal.simps(1)) lemma convex_on_ereal_sum: fixes a :: "'a \ real" fixes y :: "'a \ 'b::real_vector" fixes f :: "'b \ ereal" assumes "finite s" "s \ {}" assumes "convex_on C f" assumes "convex C" assumes "(SUM i : s. a i) = 1" assumes "\i. i \ s \ a i \ 0" assumes "\i. i \ s \ y i \ C" shows "f (SUM i : s. a i *\<^sub>R y i) \ (SUM i : s. ereal (a i) * f (y i))" using assms(1,2,5-) proof (induct s arbitrary:a rule:finite_ne_induct) case (singleton i) hence ai: "a i = 1" by auto thus ?case by (auto simp: one_ereal_def[symmetric]) next case (insert i s) from \convex_on C f\ have conv: "\x y m. ((x \ C \ y \ C \ 0 \ m \ m \ 1) \ f (m *\<^sub>R x + (1 - m) *\<^sub>R y) \ (ereal m) * f x + (1 - ereal m) * f y)" using convex_on_ereal_alt[of C f] \convex C\ by auto { assume "a i = 1" hence "(SUM j : s. a j) = 0" using insert by auto hence "\j. (j \ s \ a j = 0)" using insert by (simp add: sum_nonneg_eq_0_iff) hence ?case using insert.hyps(1-3) \a i = 1\ by (simp add: zero_ereal_def[symmetric] one_ereal_def[symmetric]) } moreover { assume asm: "a i \ 1" from insert have yai: "y i : C" "a i \ 0" by auto have fis: "finite (insert i s)" using insert by auto hence ai1: "a i \ 1" using sum_nonneg_leq_bound[of "insert i s" a] insert by simp hence "a i < 1" using asm by auto hence i0: "1 - a i > 0" by auto hence i1: "1 - ereal (a i) > 0" using ereal_minus(1)[of 1 "a i"] by (simp add: zero_ereal_def[symmetric] one_ereal_def[symmetric]) have i2: "1 - ereal (a i) \ \" using ereal_minus(1)[of 1] by (simp add: zero_ereal_def[symmetric] one_ereal_def[symmetric]) let "?a j" = "a j / (1 - a i)" have a_nonneg: "\j. j \ s \ 0 \ a j / (1 - a i)" using i0 insert by (metis insert_iff divide_nonneg_pos) have "(SUM j : insert i s. a j) = 1" using insert by auto hence "(SUM j : s. a j) = 1 - a i" using sum.insert insert by fastforce hence "(SUM j : s. a j) / (1 - a i) = 1" using i0 by auto hence a1: "(SUM j : s. ?a j) = 1" unfolding sum_divide_distrib by simp have asum: "(SUM j : s. ?a j *\<^sub>R y j) : C" using insert convex_sum[OF \finite s\ \convex C\ a1 a_nonneg] by auto have asum_le: "f (SUM j : s. ?a j *\<^sub>R y j) \ (SUM j : s. ereal (?a j) * f (y j))" using a_nonneg a1 insert by blast have "f (SUM j : insert i s. a j *\<^sub>R y j) = f ((SUM j : s. a j *\<^sub>R y j) + a i *\<^sub>R y i)" using sum.insert[of s i "\j. a j *\<^sub>R y j", OF \finite s\ \i \ s\] by (auto simp only:add.commute) also have "\ = f (((1 - a i) * inverse (1 - a i)) *\<^sub>R (SUM j : s. a j *\<^sub>R y j) + a i *\<^sub>R y i)" using i0 by auto also have "\ = f ((1 - a i) *\<^sub>R (SUM j : s. (a j * inverse (1 - a i)) *\<^sub>R y j) + a i *\<^sub>R y i)" using scaleR_right.sum[of "inverse (1 - a i)" "\j. a j *\<^sub>R y j" s, symmetric] by (auto simp:algebra_simps) also have "\ = f ((1 - a i) *\<^sub>R (SUM j : s. ?a j *\<^sub>R y j) + a i *\<^sub>R y i)" by (auto simp: divide_inverse) also have "\ \ (1 - ereal (a i)) * f ((SUM j : s. ?a j *\<^sub>R y j)) + (ereal (a i)) * f (y i)" using conv[rule_format, of "y i" "(SUM j : s. ?a j *\<^sub>R y j)" "a i"] using yai(1) asum yai(2) ai1 by (auto simp add:add.commute) also have "\ \ (1 - ereal (a i)) * (SUM j : s. ereal (?a j) * f (y j)) + (ereal (a i)) * f (y i)" using ereal_add_right_mono[OF ereal_mult_left_mono[of _ _ "1 - ereal (a i)", OF asum_le less_imp_le[OF i1]], of "(ereal (a i)) * f (y i)"] by simp also have "\ = (SUM j : s. (1 - ereal (a i)) * (ereal (?a j) * f (y j))) + (ereal (a i)) * f (y i)" unfolding ereal_pos_sum_distrib_left[of "1 - ereal (a i)" "\j. (ereal (?a j)) * f (y j)", OF less_imp_le[OF i1] i2] by auto also have "\ = (SUM j : s. (ereal (a j)) * f (y j)) + (ereal (a i)) * f (y i)" using i0 convex_on_ereal_sum_aux by auto also have "\ = (ereal (a i)) * f (y i) + (SUM j : s. (ereal (a j)) * f (y j))" by (simp add: add.commute) also have "\ = (SUM j : insert i s. (ereal (a j)) * f (y j))" using insert by auto finally have "f (SUM j : insert i s. a j *\<^sub>R y j) \ (SUM j : insert i s. (ereal (a j)) * f (y j))" by simp } ultimately show ?case by auto qed lemma sum_2: "sum u {1::nat..2} = (u 1)+(u 2)" proof- have "{1::nat..2} = {1::nat,2}" by auto thus ?thesis by auto qed lemma convex_on_ereal_iff: assumes "convex s" shows "convex_on s f \ (\k u x. (\i\{1..k::nat}. 0 \ u i \ x i : s) \ sum u {1..k} = 1 \ f (sum (\i. u i *\<^sub>R x i) {1..k} ) \ sum (\i. (ereal (u i)) * f(x i)) {1..k} )" (is "?rhs \ ?lhs") proof- { assume "?rhs" { fix k u x have zero: "~(sum u {1..0::nat} = (1::real))" by auto assume "(\i\{1..k::nat}. (0::real) \ u i \ x i \ s)" moreover assume *: "sum u {1..k} = 1" moreover from * have "k \ 0" using zero by metis ultimately have "f (sum (\i. u i *\<^sub>R x i) {1..k} ) \ sum (\i. (ereal (u i)) * f(x i)) {1..k}" using convex_on_ereal_sum[of "{1..k}" s f u x] using assms \?rhs\ by auto } hence "?lhs" by auto } moreover { assume "?lhs" { fix x y u v assume xys: "x:s" "y:s" assume uv1: "u\0" "v\0" "u + v = (1::real)" define xy where "xy = (\i::nat. if i=1 then x else y)" define uv where "uv = (\i::nat. if i=1 then u else v)" have "\i\{1..2::nat}. (0 \ uv i) \ (xy i : s)" unfolding xy_def uv_def using xys uv1 by auto moreover have "sum uv {1..2} = 1" using sum_2[of uv] unfolding uv_def using uv1 by auto moreover have "(SUM i = 1..2. uv i *\<^sub>R xy i) = u *\<^sub>R x + v *\<^sub>R y" using sum_2[of "(\i. uv i *\<^sub>R xy i)"] unfolding xy_def uv_def using xys uv1 by auto moreover have "(SUM i = 1..2. ereal (uv i) * f (xy i)) = ereal u * f x + ereal v * f y" using sum_2[of "(\i. ereal (uv i) * f (xy i))"] unfolding xy_def uv_def using xys uv1 by auto ultimately have "f (u *\<^sub>R x + v *\<^sub>R y) \ ereal u * f x + ereal v * f y" using \?lhs\[rule_format, of "2" uv xy] by auto } hence "?rhs" unfolding convex_on_def by auto } ultimately show ?thesis by blast qed lemma convex_Epigraph: assumes "convex S" shows "convex(Epigraph S f) \ convex_on S f" proof- { assume rhs: "convex(Epigraph S f)" { fix x y assume xy: "x:S" "y:S" fix u v ::real assume uv: "0 \ u" "0 \ v" "u + v = 1" have "f (u *\<^sub>R x + v *\<^sub>R y) \ ereal u * f x + ereal v * f y" proof- { assume "u=0 | v=0" hence ?thesis using uv by (auto simp: zero_ereal_def[symmetric] one_ereal_def[symmetric]) } moreover { assume "f x = \ | f y = \" hence ?thesis using uv by (auto simp: zero_ereal_def[symmetric] one_ereal_def[symmetric]) } moreover { assume a: "f x = -\ \ (f y \ \) \ ~(u=0)" from this obtain z where "f y \ ereal z" apply (cases "f y") by auto hence yz: "(y,z) : Epigraph S f" unfolding Epigraph_def using xy by auto { fix C::real have "(x, (1/u)*(C - v * z)) : Epigraph S f" unfolding Epigraph_def using a xy by auto hence "(u *\<^sub>R x + v *\<^sub>R y, C) : Epigraph S f" using rhs convexD[of "Epigraph S f" "(x, (1/u)*(C - v * z))" "(y,z)" u v] uv yz a by auto hence "(f (u *\<^sub>R x + v *\<^sub>R y) \ ereal C)" unfolding Epigraph_def by auto } hence "f (u *\<^sub>R x + v *\<^sub>R y) = -\" using ereal_bot by auto hence ?thesis by auto } moreover { assume a: "(f x \ \) \ f y = -\ \ ~(v=0)" from this obtain z where "f x \ ereal z" apply (cases "f x") by auto hence xz: "(x,z) : Epigraph S f" unfolding Epigraph_def using xy by auto { fix C::real have "(y, (1/v)*(C - u * z)) : Epigraph S f" unfolding Epigraph_def using a xy by auto hence "(u *\<^sub>R x + v *\<^sub>R y, C) : Epigraph S f" using rhs convexD[of "Epigraph S f" "(x,z)" "(y, (1/v)*(C - u * z))" u v] uv xz a by auto hence "(f (u *\<^sub>R x + v *\<^sub>R y) \ ereal C)" unfolding Epigraph_def by auto } hence "f (u *\<^sub>R x + v *\<^sub>R y) = -\" using ereal_bot by auto hence ?thesis by auto } moreover { assume a: "f x \ \ \ f x \ -\ \ f y \ \ \ f y \ -\" from this obtain fx fy where fxy: "f x = ereal fx \ f y = ereal fy" apply (cases "f x", cases "f y") by auto hence "(x, fx) : Epigraph S f \ (y, fy) : Epigraph S f" unfolding Epigraph_def using xy by auto hence "(u *\<^sub>R x + v *\<^sub>R y, u * fx + v * fy) : Epigraph S f" using rhs convexD[of "Epigraph S f" "(x,fx)" "(y,fy)" u v] uv by auto hence ?thesis unfolding Epigraph_def using fxy by auto } ultimately show ?thesis by blast qed } hence "convex_on S f" unfolding convex_on_def by auto } moreover { assume lhs: "convex_on S f" { fix x y fx fy assume xy: "(x,fx):Epigraph S f" "(y,fy):Epigraph S f" fix u v ::real assume uv: "0 \ u" "0 \ v" "u + v = 1" hence le: "f x \ ereal fx \ f y \ ereal fy" using xy unfolding Epigraph_def by auto have "x:S \ y:S" using xy unfolding Epigraph_def by auto moreover hence inS: "u *\<^sub>R x + v *\<^sub>R y : S" using assms uv convexD[of S] by auto ultimately have "f(u *\<^sub>R x + v *\<^sub>R y) \ (ereal u) * f x + (ereal v) * f y" using lhs convex_on_ereal_mem[of S f x y u v] uv by auto also have "\ \ (ereal u) * (ereal fx) + (ereal v) * (ereal fy)" apply (subst add_mono) apply (subst ereal_mult_left_mono) prefer 4 apply (subst ereal_mult_left_mono) using le uv by auto also have "\ = ereal (u * fx + v * fy)" by auto finally have "(u *\<^sub>R x + v *\<^sub>R y, u * fx + v * fy) : Epigraph S f" unfolding Epigraph_def using inS by auto } hence "convex(Epigraph S f)" unfolding convex_def by auto } ultimately show ?thesis by auto qed lemma convex_EpigraphI: "convex_on s f \ convex s \ convex(Epigraph s f)" unfolding convex_Epigraph by auto definition concave_on :: "'a::real_vector set \ ('a \ ereal) \ bool" where "concave_on S f \ convex_on S (\x. - f x)" definition finite_on :: "'a::real_vector set \ ('a \ ereal) \ bool" where "finite_on S f \ (\x\S. (f x \ \ \ f x \ -\))" definition affine_on :: "'a::real_vector set \ ('a \ ereal) \ bool" where "affine_on S f \ (convex_on S f \ concave_on S f \ finite_on S f)" definition "domain (f::_ \ ereal) = {x. f x < \}" lemma domain_Epigraph_aux: assumes "x \ \" shows "\r. x\ereal r" using assms by (cases x) auto lemma domain_Epigraph: "domain f = {x. \y. (x,y) \ Epigraph UNIV f}" unfolding domain_def Epigraph_def using domain_Epigraph_aux by auto lemma domain_Epigraph_fst: "domain f = fst ` (Epigraph UNIV f)" proof- { fix x assume "x:domain f" from this obtain y where "(x,y):Epigraph UNIV f" using domain_Epigraph[of f] by auto moreover have "x = fst (x,y)" by auto ultimately have "x:fst ` (Epigraph UNIV f)" unfolding image_def by blast } from this show ?thesis using domain_Epigraph[of f] by auto qed lemma convex_on_domain: "convex_on (domain f) f \ convex_on UNIV f" proof- { assume lhs: "convex_on (domain f) f" { fix x y fix u v ::real assume uv: "0 \ u" "0 \ v" "u + v = 1" have "f (u *\<^sub>R x + v *\<^sub>R y) \ ereal u * f x + ereal v * f y" proof- { assume "f x = \ | f y = \" hence ?thesis using uv by (auto simp: zero_ereal_def[symmetric] one_ereal_def[symmetric]) } moreover { assume "~ (f x = \ | f y = \)" hence "x : domain f \ y : domain f" unfolding domain_def by auto hence ?thesis using lhs unfolding convex_on_def using uv by auto } ultimately show ?thesis by blast qed } hence "convex_on UNIV f" unfolding convex_on_def by auto } from this show ?thesis using convex_on_ereal_subset by auto qed lemma convex_on_domain2: "convex_on (domain f) f \ (\S. convex_on S f)" by (metis convex_on_domain convex_on_ereal_univ) lemma convex_domain: fixes f :: "'a::euclidean_space \ ereal" assumes "convex_on UNIV f" shows "convex (domain f)" proof- have "convex (Epigraph UNIV f)" using assms convex_Epigraph by auto thus ?thesis unfolding domain_Epigraph_fst apply (subst convex_linear_image) using fst_linear linear_conv_bounded_linear by auto qed lemma infinite_convex_domain_iff: fixes f :: "'a::euclidean_space \ ereal" assumes "\x. (f x = \ | f x = -\)" shows "convex_on UNIV f \ convex (domain f)" proof- { assume dom: "convex (domain f)" { fix x y assume "x:domain f" "y:domain f" moreover fix u v ::real assume uv: "0 \ u" "0 \ v" "u + v = 1" ultimately have "u *\<^sub>R x + v *\<^sub>R y : domain f" using dom unfolding convex_def by auto hence "f(u *\<^sub>R x + v *\<^sub>R y) = -\" using assms unfolding domain_def by auto } hence "convex_on (domain f) f" unfolding convex_on_def by auto hence "convex_on UNIV f" by (metis convex_on_domain2) } thus ?thesis by (metis convex_domain) qed lemma convex_PInfty_outside: fixes f :: "'a::euclidean_space \ ereal" assumes "convex_on UNIV f" "convex S" shows "convex_on UNIV (\x. if x:S then (f x) else \)" proof- define g where "g x = (if x:S then -\ else \::ereal)" for x hence "convex_on UNIV g" apply (subst infinite_convex_domain_iff) using assms unfolding domain_def by auto moreover have "(\x. if x:S then (f x) else \) = (\x. max (f x) (g x))" apply (subst fun_eq_iff) unfolding g_def by auto ultimately show ?thesis using convex_ereal_max assms by auto qed definition proper_on :: "'a::real_vector set \ ('a \ ereal) \ bool" where "proper_on S f \ ((\x\S. f x \ -\) \ (\x\S. f x \ \))" definition proper :: "('a::real_vector \ ereal) \ bool" where "proper f \ proper_on UNIV f" lemma proper_iff: "proper f \ ((\x. f x \ -\) \ (\x. f x \ \))" unfolding proper_def proper_on_def by auto lemma improper_iff: "~(proper f) \ ((\x. f x = -\) | (\x. f x = \))" using proper_iff by auto lemma ereal_MInf_plus[simp]: "-\ + x = (if x = \ then \ else -\::ereal)" by (cases x) auto lemma convex_improper: fixes f :: "'a::euclidean_space \ ereal" assumes "convex_on UNIV f" assumes "~(proper f)" shows "\x\rel_interior(domain f). f x = -\" proof- { assume "domain f = {}" hence ?thesis using rel_interior_empty by auto } moreover { assume nemp: "domain f \ {}" then obtain u where u_def: "f u = -\" using assms improper_iff[of f] unfolding domain_def by auto hence udom: "u:domain f" unfolding domain_def by auto { fix x assume "x:rel_interior(domain f)" then obtain m where m_def: "m>1 \ (1 - m) *\<^sub>R u + m *\<^sub>R x : domain f" using convex_rel_interior_iff[of "domain f" x] nemp convex_domain[of f] assms udom by auto define v where "v = 1/m" hence v01: "0 v<1" using m_def by auto define y where "y = (1 - m) *\<^sub>R u + m *\<^sub>R x" have "x = (1-v) *\<^sub>R u+v *\<^sub>R y" unfolding v_def y_def using m_def by (simp add: algebra_simps) hence "f x \ (1-ereal v) * f u+ereal v * f y" using convex_on_ereal_alt_mem[of "UNIV" f y u v] assms convex_UNIV v01 by (simp add: add.commute) moreover have "f y < \" using m_def y_def unfolding domain_def by auto moreover have *: "0 < 1 - ereal v" using v01 by (metis diff_gt_0_iff_gt ereal_less(2) ereal_minus(1) one_ereal_def) moreover from * have "(1-ereal v) * f u = -\" using u_def by auto ultimately have "f x = -\" using v01 by simp } hence ?thesis by auto } ultimately show ?thesis by blast qed lemma convex_improper2: fixes f :: "'a::euclidean_space \ ereal" assumes "convex_on UNIV f" assumes "~(proper f)" shows "f x = \ | f x = -\ | x : rel_frontier (domain f)" proof- { assume a: "~(f x = \ | f x = -\)" hence "x: domain f" unfolding domain_def by auto hence "x : closure (domain f)" using closure_subset by auto moreover have "x \ rel_interior (domain f)" using assms convex_improper a by auto ultimately have "x : rel_frontier (domain f)" unfolding rel_frontier_def by auto } thus ?thesis by auto qed lemma convex_lsc_improper: fixes f :: "'a::euclidean_space \ ereal" assumes "convex_on UNIV f" assumes "~(proper f)" assumes "lsc f" shows "f x = \ | f x = -\" proof- { fix x assume "f x \ \" hence "lsc_at x f" using assms unfolding lsc_def by auto have "x: domain f" unfolding domain_def using \f x \ \\ by auto hence "x : closure (domain f)" using closure_subset by auto hence x_def: "x : closure (rel_interior (domain f))" by (metis assms(1) convex_closure_rel_interior convex_domain) { fix C assume "C0 \ (\y. dist x y < d \ C < f y)" using lst_at_delta[of x f] \lsc_at x f\ by auto from this obtain y where y_def: "y:rel_interior (domain f) \ dist y x < d" using x_def closure_approachable[of x "rel_interior (domain f)"] by auto hence "f y = -\" by (metis assms(1) assms(2) convex_improper) moreover have "C < f y" using y_def d_def by (simp add: dist_commute) ultimately have False by auto } hence "f x = -\" by auto } from this show ?thesis by auto qed lemma convex_lsc_hull: fixes f :: "'a::euclidean_space \ ereal" assumes "convex_on UNIV f" shows "convex_on UNIV (lsc_hull f)" proof- have "convex(Epigraph UNIV f)" by (metis assms convex_EpigraphI convex_UNIV) hence "convex (Epigraph UNIV (lsc_hull f))" by (metis convex_closure epigraph_lsc_hull) thus ?thesis by (metis convex_Epigraph convex_UNIV) qed lemma improper_lsc_hull: fixes f :: "'a::euclidean_space \ ereal" assumes "~(proper f)" shows "~(proper (lsc_hull f))" proof- { assume *: "\x. f x = \" then have "lsc f" by (metis (full_types) UNIV_I lsc_at_open lsc_def open_UNIV) with * have "\x. (lsc_hull f) x = \" by (metis lsc_hull_lsc) } hence "(\x. f x = \) \ (\x. (lsc_hull f) x = \)" by (metis ereal_infty_less_eq(1) lsc_hull_le) moreover have "(\x. f x = -\) \ (\x. (lsc_hull f) x = -\)" by (metis ereal_infty_less_eq2(2) lsc_hull_le) ultimately show ?thesis using assms unfolding improper_iff by auto qed lemma lsc_hull_convex_improper: fixes f :: "'a::euclidean_space \ ereal" assumes "convex_on UNIV f" assumes "~(proper f)" shows "\x\rel_interior(domain f). (lsc_hull f) x = f x" by (metis assms convex_improper ereal_infty_less_eq(2) lsc_hull_le) lemma convex_with_rel_open_domain: fixes f :: "'a::euclidean_space \ ereal" assumes "convex_on UNIV f" assumes "rel_open (domain f)" shows "(\x. f x > -\) | (\x. (f x = \ | f x = -\))" proof- { assume "\(\x. f x > -\)" hence "\(proper f)" using proper_iff by auto hence "\x\rel_interior(domain f). f x = -\" by (metis assms(1) convex_improper) hence "\x\domain f. f x = -\" by (metis assms(2) rel_open_def) hence "\x. (f x = \ | f x = -\)" unfolding domain_def by auto } thus ?thesis by auto qed lemma convex_with_UNIV_domain: fixes f :: "'a::euclidean_space \ ereal" assumes "convex_on UNIV f" assumes "domain f = UNIV" shows "(\x. f x > -\) \ (\x. f x = -\)" by (metis assms convex_improper ereal_MInfty_lessI proper_iff rel_interior_UNIV UNIV_I) lemma rel_interior_Epigraph: fixes f :: "'a::euclidean_space \ ereal" assumes "convex_on UNIV f" shows "(x,z) : rel_interior (Epigraph UNIV f) \ (x : rel_interior (domain f) \ f x < ereal z)" apply (subst rel_interior_projection[of _ "(\y. {z. (y, z) : Epigraph UNIV f})"]) apply (metis assms convex_EpigraphI convex_UNIV convex_on_ereal_univ) unfolding domain_Epigraph Epigraph_def using rel_interior_ereal_semiline by auto lemma rel_interior_EpigraphI: fixes f :: "'a::euclidean_space \ ereal" assumes "convex_on UNIV f" shows "rel_interior (Epigraph UNIV f) = {(x,z) |x z. x : rel_interior (domain f) \ f x < ereal z}" proof- { fix x z have "(x,z):rel_interior (Epigraph UNIV f) \ (x : rel_interior (domain f) \ f x < ereal z)" using rel_interior_Epigraph[of f x z] assms by auto } thus ?thesis by auto qed lemma convex_less_ri_domain: fixes f :: "'a::euclidean_space \ ereal" assumes "convex_on UNIV f" assumes "\x. f x < a" shows "\x\rel_interior (domain f). f x < a" proof- define A where "A = {(x::'a::euclidean_space,m)|x m. ereal m ereal z < a" using ereal_dense2 by auto hence "(x,z):A \ (x,z):Epigraph UNIV f" unfolding A_def Epigraph_def by auto hence "A Int (Epigraph UNIV f) \ {}" unfolding A_def Epigraph_def using assms by auto moreover have "open A" proof(cases a) case real hence "A = {y. inner (0::'a, 1) y < real_of_ereal a}" using A_def by auto thus ?thesis using open_halfspace_lt by auto next case PInf thus ?thesis using A_def by auto next case MInf thus ?thesis using A_def by auto qed ultimately have "A Int rel_interior(Epigraph UNIV f) \ {}" by (metis assms(1) convex_Epigraph convex_UNIV open_Int_closure_eq_empty open_inter_closure_rel_interior) then obtain x0 z0 where "(x0,z0):A \ x0 : rel_interior (domain f) \ f x0 < ereal z0" using rel_interior_Epigraph[of f] assms by auto thus ?thesis apply(rule_tac x="x0" in bexI) using A_def by auto qed lemma rel_interior_eq_between: fixes S T :: "('m::euclidean_space) set" assumes "convex S" "convex T" shows "(rel_interior S = rel_interior T) \ (rel_interior S \ T \ T \ closure S)" by (metis assms closure_eq_between convex_closure_rel_interior convex_rel_interior_closure) lemma convex_less_in_riS: fixes f :: "'a::euclidean_space \ ereal" assumes "convex_on UNIV f" assumes "convex S" "rel_interior S \ domain f" assumes "\x\closure S. f x < a" shows "\x\rel_interior S. f x < a" proof- define g where "g x = (if x:closure S then (f x) else \)" for x hence "\x. g x < a" using assms by auto have convg: "convex_on UNIV g" unfolding g_def apply (subst convex_PInfty_outside) using assms convex_closure by auto hence *: "\x\rel_interior (domain g). g x < a" apply (subst convex_less_ri_domain) using assms g_def by auto have "convex (domain g)" by (metis convg convex_domain) moreover have "rel_interior S \ domain g \ domain g \ closure S" using g_def assms rel_interior_subset_closure unfolding domain_def by auto ultimately have "rel_interior (domain g) = rel_interior S" by (metis assms(2) rel_interior_eq_between) thus ?thesis by (metis "*" g_def less_ereal.simps(2)) qed lemma convex_less_inS: fixes f :: "'a::euclidean_space \ ereal" assumes "convex_on UNIV f" assumes "convex S" "S \ domain f" assumes "\x\closure S. f x < a" shows "\x\S. f x < a" using convex_less_in_riS[of f S a] rel_interior_subset[of S] assms by auto lemma convex_finite_geq_on_closure: fixes f :: "'a::euclidean_space \ ereal" assumes "convex_on UNIV f" assumes "convex S" "finite_on S f" assumes "\x\S. f x \ a" shows "\x\closure S. f x \ a" proof- have "S \ domain f" using assms unfolding finite_on_def domain_def by auto { assume "\(\x\closure S. f x \ a)" hence "\x\closure S. f x < a" by (simp add: not_le) hence False using convex_less_inS[of f S a] assms \S \ domain f\ by auto } thus ?thesis by auto qed lemma lsc_hull_of_convex_determined: fixes f g :: "'a::euclidean_space \ ereal" assumes "convex_on UNIV f" "convex_on UNIV g" assumes "rel_interior (domain f) = rel_interior (domain g)" assumes "\x\rel_interior (domain f). f x = g x" shows "lsc_hull f = lsc_hull g" proof- have "rel_interior (Epigraph UNIV f) = rel_interior (Epigraph UNIV g)" apply (subst rel_interior_EpigraphI, metis assms)+ using assms by auto hence "closure (Epigraph UNIV f) = closure (Epigraph UNIV g)" by (metis assms convex_EpigraphI convex_UNIV convex_closure_rel_interior) thus ?thesis by (metis lsc_hull_expl) qed lemma domain_lsc_hull_between: fixes f :: "'a::euclidean_space \ ereal" shows "domain f \ domain (lsc_hull f) \ domain (lsc_hull f) \ closure (domain f)" proof- { fix x assume "x:domain f" hence "x:domain (lsc_hull f)" unfolding domain_def using lsc_hull_le[of f x] by auto } moreover { fix x assume "x:domain (lsc_hull f)" hence "min (f x) (Liminf (at x) f) < \" unfolding domain_def using lsc_hull_liminf_at[of f] by auto then obtain z where z_def: "min (f x) (Liminf (at x) f) < z \ z < \" by (metis dense) { fix e::real assume "e>0" hence "Inf (f ` ball x e) \ min (f x) (Liminf (at x) f)" unfolding min_Liminf_at apply (subst SUP_upper) by auto hence "\y. y \ ball x e \ f y \ z" using Inf_le_iff_less [of f "ball x e" "min (f x) (Liminf (at x) f)"] z_def by (auto simp add: Bex_def) hence "\y. dist x y < e \ y \ domain f" unfolding domain_def ball_def using z_def by auto } hence "x\closure(domain f)" unfolding closure_approachable by (auto simp add: dist_commute) } ultimately show ?thesis by auto qed lemma domain_vs_domain_lsc_hull: fixes f :: "'a::euclidean_space \ ereal" assumes "convex_on UNIV f" shows "rel_interior(domain (lsc_hull f)) = rel_interior(domain f) \ closure(domain (lsc_hull f)) = closure(domain f) \ aff_dim(domain (lsc_hull f)) = aff_dim(domain f)" proof- have "convex (domain f)" by (metis assms convex_domain) moreover have "convex (domain (lsc_hull f))" by (metis assms convex_domain convex_lsc_hull) moreover have "rel_interior (domain f) \ domain (lsc_hull f) \ domain (lsc_hull f) \ closure (domain f)" by (metis domain_lsc_hull_between rel_interior_subset subset_trans) ultimately show ?thesis by (metis closure_eq_between rel_interior_aff_dim rel_interior_eq_between) qed lemma vertical_line_affine: fixes x :: "'a::euclidean_space" shows "affine {(x,m::real)|m. m:UNIV}" unfolding affine_def by (auto simp add: pth_8) lemma lsc_hull_of_convex_agrees_onRI: fixes f :: "'a::euclidean_space \ ereal" assumes "convex_on UNIV f" shows "\x\rel_interior (domain f). (f x = (lsc_hull f) x)" proof- have cEpi: "convex (Epigraph UNIV f)" by (metis assms convex_EpigraphI convex_UNIV) { fix x assume x_def: "x : rel_interior (domain f)" hence "f x < \" unfolding domain_def using rel_interior_subset by auto then obtain r where r_def: "(x,r):rel_interior (Epigraph UNIV f)" using assms x_def rel_interior_Epigraph[of f x] by (metis ereal_dense2) define M where "M = {(x,m::real)|m. m:UNIV}" hence "affine M" using vertical_line_affine by auto moreover have "rel_interior (Epigraph UNIV f) Int M \ {}" using r_def M_def by auto ultimately have *: "closure (Epigraph UNIV f) Int M = closure (Epigraph UNIV f Int M)" using convex_affine_closure_Int[of "Epigraph UNIV f" M] cEpi by auto have "Epigraph UNIV f Int M = {x} \ {m. f x \ ereal m}" unfolding Epigraph_def M_def by auto moreover have "closed({x} \ {m. f x\ ereal m})" apply (subst closed_Times) using closed_ereal_semiline by auto ultimately have "{x} \ {m. f x \ ereal m} = closure (Epigraph UNIV f) Int M" by (metis "*" Int_commute closure_closed) also have "\=Epigraph UNIV (lsc_hull f) Int M" by (metis Int_commute epigraph_lsc_hull) also have "\={x} \ {m. ((lsc_hull f) x) \ ereal m}" unfolding Epigraph_def M_def by auto finally have "{m. f x \ ereal m} = {m. lsc_hull f x \ ereal m}" by auto hence "f x = (lsc_hull f) x" using ereal_semiline_unique by auto } thus ?thesis by auto qed lemma lsc_hull_of_convex_agrees_outside: fixes f :: "'a::euclidean_space \ ereal" assumes "convex_on UNIV f" shows "\x. x \ closure (domain f) \ (f x = (lsc_hull f) x)" proof- { fix x assume "x \ closure (domain f)" hence "x \ domain (lsc_hull f)" using domain_lsc_hull_between by auto hence "(lsc_hull f) x = \" unfolding domain_def by auto hence "f x = (lsc_hull f) x" using lsc_hull_le[of f x] by auto } thus ?thesis by auto qed lemma lsc_hull_of_convex_agrees: fixes f :: "'a::euclidean_space \ ereal" assumes "convex_on UNIV f" shows "\x. (f x = (lsc_hull f) x) | x : rel_frontier (domain f)" by (metis DiffI assms lsc_hull_of_convex_agrees_onRI lsc_hull_of_convex_agrees_outside rel_frontier_def) lemma lsc_hull_of_proper_convex_proper: fixes f :: "'a::euclidean_space \ ereal" assumes "convex_on UNIV f" "proper f" shows "proper (lsc_hull f)" proof- obtain x where x_def: "x:rel_interior (domain f) \ f x < \" by (metis assms convex_less_ri_domain ereal_less_PInfty proper_iff) hence "f x = (lsc_hull f) x" using lsc_hull_of_convex_agrees[of f] assms unfolding rel_frontier_def by auto moreover have "f x > -\" using assms proper_iff by auto ultimately have "(lsc_hull f) x < \ \ (lsc_hull f) x > -\" using x_def by auto thus ?thesis using convex_lsc_improper[of "lsc_hull f" x] lsc_lsc_hull[of f] assms convex_lsc_hull[of f] by auto qed lemma lsc_hull_of_proper_convex: fixes f :: "'a::euclidean_space \ ereal" assumes "convex_on UNIV f" "proper f" shows "lsc (lsc_hull f) \ proper (lsc_hull f) \ convex_on UNIV (lsc_hull f) \ (\x. (f x = (lsc_hull f) x) | x : rel_frontier (domain f))" by (metis assms convex_lsc_hull lsc_hull_of_convex_agrees lsc_hull_of_proper_convex_proper lsc_lsc_hull) lemma affine_no_rel_frontier: fixes S :: "('n::euclidean_space) set" assumes "affine S" shows "rel_frontier S = {}" unfolding rel_frontier_def using assms affine_closed[of S] closure_closed[of S] affine_rel_open[of S] rel_open_def[of S] by auto lemma convex_with_affine_domain_is_lsc: fixes f :: "'a::euclidean_space \ ereal" assumes "convex_on UNIV f" assumes "affine (domain f)" shows "lsc f" by (metis assms affine_no_rel_frontier emptyE lsc_def lsc_hull_liminf_at lsc_hull_of_convex_agrees lsc_liminf_at_eq) lemma convex_finite_is_lsc: fixes f :: "'a::euclidean_space \ ereal" assumes "convex_on UNIV f" assumes "finite_on UNIV f" shows "lsc f" proof- have "affine (domain f)" using assms affine_UNIV unfolding finite_on_def domain_def by auto thus ?thesis by (metis assms(1) convex_with_affine_domain_is_lsc) qed lemma always_eventually_within: "(\x\S. P x) \ eventually P (at x within S)" unfolding eventually_at_filter by auto lemma ereal_divide_pos: assumes "(a::ereal)>0" "b>0" shows "a/(ereal b)>0" by (metis PInfty_eq_infinity assms ereal.simps(2) ereal_less(2) ereal_less_divide_pos ereal_mult_zero) lemma real_interval_limpt: assumes "a0 \ cball b e \ T" using open_contains_cball[of T] by auto hence "(b-e):cball b e" unfolding cball_def dist_norm by auto moreover { assume "a\b-e" hence "a:cball b e" unfolding cball_def dist_norm using \a by auto } ultimately have "max a (b-e):cball b e" by (metis max.absorb1 max.absorb2 linear) hence "max a (b-e):T" using e_def by auto moreover have "max a (b-e):{a..a by auto ultimately have "\y\{a.. y \ b" by auto } thus ?thesis unfolding islimpt_def by auto qed lemma lsc_hull_of_convex_aux: "Limsup (at 1 within {0..<1}) (\m. ereal ((1-m)*a+m*b)) \ ereal b" proof- have nontr: "~trivial_limit (at 1 within {0..< 1::real})" apply (subst trivial_limit_within) using real_interval_limpt by auto have "((\m. ereal ((1-m)*a+m*b)) \ (1 - 1) * a + 1 * b) (at 1 within {0..<1})" unfolding lim_ereal by (intro tendsto_intros) from lim_imp_Limsup[OF nontr this] show ?thesis by simp qed lemma lsc_hull_of_convex: fixes f :: "'a::euclidean_space \ ereal" assumes "convex_on UNIV f" assumes "x : rel_interior (domain f)" shows "((\m. f((1-m)*\<^sub>R x + m *\<^sub>R y)) \ (lsc_hull f) y) (at 1 within {0..<1})" (is "(?g \ _ y) _") proof (cases "y=x") case True hence "?g = (\m. f y)" by (simp add: algebra_simps) hence "(?g \ f y) (at 1 within {0..<1})" by simp moreover have "(lsc_hull f) y = f y" by (metis \y=x\ assms lsc_hull_of_convex_agrees_onRI) ultimately show ?thesis by auto next case False have aux: "\m. y - ((1 - m) *\<^sub>R x + m *\<^sub>R y) = (1-m)*\<^sub>R (y-x)" by (simp add: algebra_simps) have "(lsc_hull f) y = min (f y) (Liminf (at y) f)" by (metis lsc_hull_liminf_at) also have "\ \ Liminf (at 1 within {0..<1}) ?g" unfolding min_Liminf_at unfolding Liminf_within apply (subst SUP_mono) apply (rule_tac x="n/norm(y-x)" in bexI) apply (subst INF_mono) apply (rule_tac x="(1 - m) *\<^sub>R x + m *\<^sub>R y" in bexI) prefer 2 unfolding ball_def dist_norm by (auto simp add: aux \y\x\ less_divide_eq) finally have *: "(lsc_hull f) y \ Liminf (at 1 within {0..<1}) ?g" by auto { fix b assume "ereal b\(lsc_hull f) y" hence yb: "(y,b):closure(Epigraph UNIV f)" by (metis epigraph_lsc_hull mem_Epigraph UNIV_I) have "x : domain f" by (metis assms(2) rel_interior_subset rev_subsetD) hence "f x<\" unfolding domain_def by auto then obtain a where"ereal a>f x" by (metis ereal_dense2) hence xa: "(x,a):rel_interior(Epigraph UNIV f)" by (metis assms rel_interior_Epigraph) { fix m :: real assume "0 \ m \ m < 1" hence "(y, b) - (1-m) *\<^sub>R ((y, b) - (x, a)) : rel_interior (Epigraph UNIV f)" apply (subst rel_interior_closure_convex_shrink) apply (metis assms(1) convex_Epigraph convex_UNIV convex_on_ereal_univ) using yb xa by auto hence "f (y - (1-m) *\<^sub>R (y-x) ) < ereal (b-(1-m)*(b-a))" using assms(1) rel_interior_Epigraph by auto hence "?g m \ ereal ((1-m)*a+m*b)" by (simp add: algebra_simps) } hence "eventually (\m. ?g m \ ereal ((1-m)*a+m*b)) (at 1 within {0..<1})" apply (subst always_eventually_within) by auto hence "Limsup (at 1 within {0..<1}) ?g \ Limsup (at 1 within {0..<1}) (\m. ereal ((1-m)*a+m*b))" apply (subst Limsup_mono) by auto also have "\ \ ereal b" using lsc_hull_of_convex_aux by auto finally have "Limsup (at 1 within {0..<1}) ?g \ereal b" by auto } hence "Limsup (at 1 within {0..<1}) ?g \ (lsc_hull f) y" using ereal_le_real[of "(lsc_hull f) y"] by auto moreover have nontr: "~trivial_limit (at (1::real) within {0..<1})" apply (subst trivial_limit_within) using real_interval_limpt by auto moreover hence "Liminf (at 1 within {0..<1}) ?g \ Limsup (at 1 within {0..<1}) ?g" apply (subst Liminf_le_Limsup) by auto ultimately have "Limsup (at 1 within {0..<1}) ?g = (lsc_hull f) y \ Liminf (at 1 within {0..<1}) ?g = (lsc_hull f) y" using * by auto thus ?thesis apply (subst Liminf_eq_Limsup) using nontr by auto qed end diff --git a/thys/Octonions/Cross_Product_7.thy b/thys/Octonions/Cross_Product_7.thy --- a/thys/Octonions/Cross_Product_7.thy +++ b/thys/Octonions/Cross_Product_7.thy @@ -1,379 +1,378 @@ (* Title: Cross_Product_7.thy Author: Angeliki Koutsoukou-Argyraki, University of Cambridge Date: September 2018 *) section\Vector Cross Product in 7 Dimensions\ theory Cross_Product_7 - imports "HOL-Analysis.Ordered_Euclidean_Space" - "HOL-Analysis.Determinants" + imports "HOL-Analysis.Multivariate_Analysis" begin subsection\Elementary auxiliary lemmas.\ lemma exhaust_7: fixes x :: 7 shows "x = 1 \ x = 2 \ x = 3 \ x = 4 \ x = 5 \ x = 6 \ x = 7 " proof (induct x) case (of_int z) then have "0 \ z" and "z < 7" by simp_all then have "z = 0 \ z = 1 \ z = 2 \ z = 3 \ z = 4 \ z =5\ z = 6 \ z =7 " by arith then show ?case by auto qed lemma forall_7: "(\i::7. P i) \ P 1 \ P 2 \ P 3\ P 4 \ P 5 \ P 6\ P 7 " by (metis exhaust_7) lemma vector_7 [simp]: "(vector [x1,x2,x3,x4,x5,x6,x7] ::('a::zero)^7)$1 = x1" "(vector [x1,x2,x3,x4,x5,x6,x7] ::('a::zero)^7)$2 = x2" "(vector [x1,x2,x3,x4,x5,x6,x7] ::('a::zero)^7)$3 = x3" "(vector [x1,x2,x3,x4,x5,x6,x7] ::('a::zero)^7)$4 = x4" "(vector [x1,x2,x3,x4,x5,x6,x7] ::('a::zero)^7)$5 = x5" "(vector [x1,x2,x3,x4,x5,x6,x7] ::('a::zero)^7)$6 = x6" "(vector [x1,x2,x3,x4,x5,x6,x7] ::('a::zero)^7)$7 = x7" unfolding vector_def by auto lemma forall_vector_7: "(\v::'a::zero^7. P v) \ (\x1 x2 x3 x4 x5 x6 x7. P(vector[x1, x2, x3, x4, x5, x6, x7]))" apply auto apply (erule_tac x="v$1" in allE) apply (erule_tac x="v$2" in allE) apply (erule_tac x="v$3" in allE) apply (erule_tac x="v$4" in allE) apply (erule_tac x="v$5" in allE) apply (erule_tac x="v$6" in allE) apply (erule_tac x="v$7" in allE) apply (subgoal_tac "vector [v$1, v$2, v$3, v$4, v$5, v$6, v$7] = v") apply simp apply (vector vector_def) apply (simp add: forall_7) done lemma UNIV_7: "UNIV = {1::7, 2::7, 3::7, 4::7, 5::7, 6::7, 7::7}" using exhaust_7 by auto lemma sum_7: "sum f (UNIV::7 set) = f 1 + f 2 + f 3 + f 4 + f 5 + f 6 + f 7" unfolding UNIV_7 by (simp add: ac_simps) lemma not_equal_vector7 : fixes x::"real^7" and y::"real^7" assumes "x = vector[x1,x2,x3,x4,x5,x6,x7] " and "y= vector [y1,y2,y3,y4,y5,y6,y7]" and "x$1 \ y$1 \ x$2 \ y$2 \ x$3 \ y$3 \ x$4 \ y$4 \ x$5 \ y$5 \ x$6 \ y$6 \ x$7 \ y$7 " shows "x \ y" using assms by blast lemma equal_vector7: fixes x::"real^7" and y::"real^7" assumes "x = vector[x1,x2,x3,x4,x5,x6,x7] " and "y= vector [y1,y2,y3,y4,y5,y6,y7]" and "x = y" shows "x$1 = y$1 \ x$2 = y$2 \ x$3 = y$3 \ x$4 = y$4 \ x$5 = y$5 \ x$6 = y$6 \ x$7 = y$7 " using assms by blast lemma numeral_4_eq_4: "4 = Suc( Suc (Suc (Suc 0)))" by (simp add: eval_nat_numeral) lemma numeral_5_eq_5: "5 = Suc(Suc( Suc (Suc (Suc 0))))" by (simp add: eval_nat_numeral) lemma numeral_6_eq_6: "6 = Suc( Suc(Suc( Suc (Suc (Suc 0)))))" by (simp add: eval_nat_numeral) lemma numeral_7_eq_7: "7 = Suc(Suc( Suc(Suc( Suc (Suc (Suc 0))))))" by (simp add: eval_nat_numeral) subsection\The definition of the 7D cross product and related lemmas\ context includes no_Set_Product_syntax begin (* locally disable syntax for set product, to avoid warnings *) text \Note: in total there exist 480 equivalent multiplication tables for the definition, the following is based on the one most widely used: \ definition cross7 :: "[real^7, real^7] \ real^7" (infixr "\\<^sub>7" 80) where "a \\<^sub>7 b \ vector [a$2 * b$4 - a$4 * b$2 + a$3 * b$7 - a$7 * b$3 + a$5 * b$6 - a$6 * b$5 , a$3 * b$5 - a$5 * b$3 + a$4 * b$1 - a$1 * b$4 + a$6 * b$7 - a$7 * b$6 , a$4 * b$6 - a$6 * b$4 + a$5 * b$2 - a$2 * b$5 + a$7 * b$1 - a$1 * b$7 , a$5 * b$7 - a$7 * b$5 + a$6 * b$3 - a$3 * b$6 + a$1 * b$2 - a$2 * b$1 , a$6 * b$1 - a$1 * b$6 + a$7 * b$4 - a$4 * b$7 + a$2 * b$3 - a$3 * b$2 , a$7 * b$2 - a$2 * b$7 + a$1 * b$5 - a$5 * b$1 + a$3 * b$4 - a$4 * b$3 , a$1 * b$3 - a$3 * b$1 + a$2 * b$6 - a$6 * b$2 + a$4 * b$5 - a$5 * b$4 ]" end bundle cross7_syntax begin notation cross7 (infixr "\\<^sub>7" 80) no_notation Product_Type.Times (infixr "\\<^sub>7" 80) end bundle no_cross7_syntax begin no_notation cross7 (infixr "\\<^sub>7" 80) notation Product_Type.Times (infixr "\\<^sub>7" 80) end unbundle cross7_syntax lemmas cross7_simps = cross7_def inner_vec_def sum_7 det_def vec_eq_iff vector_def algebra_simps lemma dot_cross7_self: "x \ (x \\<^sub>7 y) = 0" "x \ (y \\<^sub>7 x) = 0" "(x \\<^sub>7 y) \ y = 0" "(y \\<^sub>7 x) \ y = 0" by (simp_all add: orthogonal_def cross7_simps) lemma orthogonal_cross7: "orthogonal (x \\<^sub>7 y) x" "orthogonal (x \\<^sub>7 y) y" "orthogonal y (x\\<^sub>7 y)" "orthogonal (x \\<^sub>7 y) x" by (simp_all add: orthogonal_def dot_cross7_self) lemma cross7_zero_left [simp]: "0 \\<^sub>7 x = 0" and cross7_zero_right [simp]: "x \\<^sub>7 0 = 0" by (simp_all add: cross7_simps) lemma cross7_skew: "(x \\<^sub>7 y) = -(y \\<^sub>7 x)" by (simp add: cross7_simps) lemma cross7_refl [simp]: "x \\<^sub>7 x = 0" by (simp add: cross7_simps) lemma cross7_add_left: "(x + y) \\<^sub>7 z = (x \\<^sub>7 z) + (y \\<^sub>7 z)" and cross7_add_right: "x \\<^sub>7 (y + z) = (x \\<^sub>7 y) + (x \\<^sub>7 z)" by (simp_all add: cross7_simps) lemma cross7_mult_left: "(c *\<^sub>R x) \\<^sub>7 y = c *\<^sub>R (x \\<^sub>7 y)" and cross7_mult_right: "x \\<^sub>7 (c *\<^sub>R y) = c *\<^sub>R (x \\<^sub>7 y)" by (simp_all add: cross7_simps) lemma cross7_minus_left [simp]: "(-x) \\<^sub>7 y = - (x \\<^sub>7 y)" and cross7_minus_right [simp]: "x \\<^sub>7 -y = - (x \\<^sub>7 y)" by (simp_all add: cross7_simps) lemma left_diff_distrib: "(x - y) \\<^sub>7 z = x \\<^sub>7 z - y \\<^sub>7 z" and right_diff_distrib: "x \\<^sub>7 (y - z) = x \\<^sub>7 y - x \\<^sub>7 z" by (simp_all add: cross7_simps) hide_fact (open) left_diff_distrib right_diff_distrib lemma cross7_triple1: "(x \\<^sub>7 y) \ z = (y \\<^sub>7 z) \ x" and cross7_triple2: "(x \\<^sub>7 y) \ z = x \ (y \\<^sub>7 z) " by (simp_all add: cross7_def inner_vec_def sum_7 vec_eq_iff algebra_simps) lemma scalar7_triple1: "x \ (y \\<^sub>7 z) = y \ (z \\<^sub>7 x)" and scalar7_triple2: "x \ (y \\<^sub>7 z) = z \ (x \\<^sub>7 y ) " by (simp_all add: cross7_def inner_vec_def sum_7 vec_eq_iff algebra_simps) lemma cross7_components: "(x \\<^sub>7 y)$1 = x$2 * y$4 - x$4 * y$2 + x$3 * y$7 - x$7 * y$3 + x$5 * y$6 - x$6 * y$5 " "(x \\<^sub>7 y)$2 = x$4 * y$1 - x$1 * y$4 + x$3 * y$5 - x$5 * y$3 + x$6 * y$7 - x$7 * y$6 " "(x \\<^sub>7 y)$3 = x$5 * y$2 - x$2 * y$5 + x$4 * y$6 - x$6 * y$4 + x$7 * y$1 - x$1 * y$7 " "(x \\<^sub>7 y)$4 = x$1 * y$2 - x$2 * y$1 + x$6 * y$3 - x$3 * y$6 + x$5 * y$7 - x$7 * y$5 " "(x \\<^sub>7 y)$5 = x$6 * y$1 - x$1 * y$6 + x$2 * y$3 - x$3 * y$2 + x$7 * y$4 - x$4 * y$7 " "(x \\<^sub>7 y)$6 = x$1 * y$5 - x$5 * y$1 + x$7 * y$2 - x$2 * y$7 + x$3 * y$4 - x$4 * y$3 " "(x \\<^sub>7 y)$7 = x$1 * y$3 - x$3 * y$1 + x$4 * y$5 - x$5 * y$4 + x$2 * y$6 - x$6 * y$2 " by (simp_all add: cross7_def inner_vec_def sum_7 vec_eq_iff algebra_simps) text \Nonassociativity of the 7D cross product showed using a counterexample\ lemma cross7_nonassociative: " \ (\ (c::real^7) (a::real^7) ( b::real^7) . c \\<^sub>7 (a \\<^sub>7 b) = (c \\<^sub>7 a ) \\<^sub>7 b) " proof- have *: " \ (c::real^7) (a::real^7) ( b::real^7) . c \\<^sub>7 (a \\<^sub>7 b) \ (c \\<^sub>7 a ) \\<^sub>7 b " proof- define f::"real^7" where "f = vector[ 0, 0, 0, 0, 0, 1, 1 ]" define g::"real^7" where "g = vector[ 0, 0, 0, 1, 0, 0, 0 ]" define h::"real^7" where "h = vector[ 1, 0, 1, 0, 0, 0, 0 ]" define u where " u= (g \\<^sub>7 h) " define v where " v= (f \\<^sub>7 g) " have 1:" u = vector[0, 1, 0, 0, 0, -1, 0]" unfolding cross7_def g_def h_def f_def u_def by simp have 3:" f \\<^sub>7 u = vector[0, 1, 0, 0, 0, 1, -1 ] " unfolding cross7_def f_def using 1 by simp have 2:" v = vector[0, 0, -1, 0, 1, 0, 0]" unfolding cross7_def g_def h_def f_def v_def by simp have 4:" v \\<^sub>7 h = vector[0, -1, 0, 0, 0, -1, 1 ] " unfolding cross7_def h_def using 2 by simp define x::"real^7" where "x= vector[0, 1, 0, 0, 0, 1, -1] " define y::"real^7" where "y= vector[0, -1, 0, 0, 0,-1, 1] " have *: "x$2 = 1" unfolding x_def by simp have **: "y$2 = -1" unfolding y_def by simp have ***: "x \ y" using * ** by auto have 5: " f \\<^sub>7 u \ v \\<^sub>7 h" unfolding x_def y_def using *** by (simp add: "3" "4" x_def y_def) have 6:" f \\<^sub>7 (g \\<^sub>7 h) \ (f \\<^sub>7 g ) \\<^sub>7 h " using 5 by (simp add: u_def v_def) show ?thesis unfolding f_def g_def h_def using 6 by blast qed show ?thesis using * by blast qed text \The 7D cross product does not satisfy the Jacobi Identity(shown using a counterexample) \ lemma cross7_not_Jacobi: " \ (\ (c::real^7) (a::real^7) ( b::real^7) . (c \\<^sub>7 a ) \\<^sub>7 b + (b \\<^sub>7 c) \\<^sub>7 a + (a \\<^sub>7 b ) \\<^sub>7 c =0 ) " proof- have *: " \ (c::real^7) (a::real^7) ( b::real^7) . (c \\<^sub>7 a ) \\<^sub>7 b + (b \\<^sub>7 c) \\<^sub>7 a + (a \\<^sub>7 b ) \\<^sub>7 c \0 " proof- define f::"real^7" where " f= vector[ 0, 0, 0, 0, 0, 1, 1 ] " define g::"real^7" where " g= vector[ 0, 0, 0, 1, 0, 0, 0 ] " define h::"real^7" where " h= vector[ 1, 0, 1, 0, 0, 0, 0 ] " define u where " u= (g \\<^sub>7 h) " define v where " v= (f \\<^sub>7 g) " define w where " w = (h \\<^sub>7 f) " have 1:" u = vector[0, 1, 0, 0, 0, -1, 0]" unfolding cross7_def g_def h_def f_def u_def by simp have 3:" u \\<^sub>7 f = vector[0,- 1, 0, 0, 0,- 1, 1 ] " unfolding cross7_def f_def using 1 by simp have 2:" v = vector[0, 0, -1, 0, 1, 0, 0]" unfolding cross7_def g_def h_def f_def v_def by simp have 4:" v \\<^sub>7 h = vector[0, -1, 0, 0, 0, -1, 1 ] " unfolding cross7_def h_def using 2 by simp have 8: " w = vector[1, 0, -1, -1, -1, 0, 0]" unfolding cross7_def h_def f_def w_def by simp have 9: " w \\<^sub>7 g = vector[0, -1, 0, 0, 0, -1, 1]" unfolding cross7_def h_def f_def w_def g_def apply simp done have &: "(u \\<^sub>7 f)$2+( v \\<^sub>7 h) $2+( w \\<^sub>7 g) $2 =-3" using 3 4 9 by simp have &&: " u \\<^sub>7 f + v \\<^sub>7 h + w \\<^sub>7 g \ 0 " using & by (metis vector_add_component zero_index zero_neq_neg_numeral) show ?thesis using && u_def v_def w_def by blast qed show ?thesis using * by blast qed text\The vector triple product property fulfilled for the 3D cross product does not hold for the 7D cross product. Shown below with a counterexample. \ lemma cross7_not_vectortriple: "\(\ (c::real^7) (a::real^7) ( b::real^7).( c \\<^sub>7 a ) \\<^sub>7 b = (c \ b ) *\<^sub>R a - (c \ a )*\<^sub>R b)" proof- have *: "\(c::real^7) (a::real^7) ( b::real^7). (c \\<^sub>7 a) \\<^sub>7 b \ (c \ b ) *\<^sub>R a - (c \ a )*\<^sub>R b" proof- define g:: "real ^ 7" where "g = vector[0, 0, 0, 1, 0, 0, 0]" define h:: "real ^ 7" where "h = vector[1, 0, 1, 0, 0, 0, 0]" define f:: "real ^ 7" where "f = vector[0, 0, 0, 0, 0, 1, 1]" define u where "u = (g \\<^sub>7 h)" have 1: "u = vector[0, 1, 0, 0, 0, -1, 0]" unfolding cross7_def g_def h_def f_def u_def by simp have 2:" u \\<^sub>7 f = vector[0,- 1, 0, 0, 0,- 1, 1 ]" unfolding cross7_def f_def using 1 by simp have 3: "(g \ f) *\<^sub>R h = 0" unfolding g_def f_def inner_vec_def by (simp add: sum_7) have 4: "(g \ h) *\<^sub>R f = 0" unfolding g_def h_def inner_vec_def by (simp add: sum_7) have 5: "(g \ f) *\<^sub>R h -(g \ h) *\<^sub>R f = 0" using 3 4 by auto have 6: "u \\<^sub>7 f \ 0" using 2 by (metis one_neq_zero vector_7(7) zero_index) have 7: "(g \\<^sub>7 h) \\<^sub>7 f \ 0" using 2 6 u_def by blast have 8: "(g \\<^sub>7 h) \\<^sub>7 f \ (g \ f) *\<^sub>R h - (g \ h ) *\<^sub>R f" using 5 7 by simp show ?thesis using 8 by auto qed show ?thesis using * by auto qed lemma axis_nth_neq [simp]: "i \ j \ axis i x $ j = 0" by (simp add: axis_def) lemma cross7_basis: "(axis 1 1) \\<^sub>7 (axis 2 1) = axis 4 1" "(axis 2 1) \\<^sub>7 (axis 1 1) = -(axis 4 1)" "(axis 2 1) \\<^sub>7 (axis 3 1) = axis 5 1" "(axis 3 1) \\<^sub>7 (axis 2 1) = -(axis 5 1)" "(axis 3 1) \\<^sub>7 (axis 4 1) = axis 6 1" "(axis 4 1) \\<^sub>7 (axis 3 1) = -(axis 6 1)" "(axis 2 1) \\<^sub>7 (axis 4 1) = axis 1 1" "(axis 4 1) \\<^sub>7 (axis 2 1) = -(axis 1 1)" "(axis 4 1) \\<^sub>7 (axis 5 1) = axis 7 1" "(axis 5 1) \\<^sub>7 (axis 4 1) = -(axis 7 1)" "(axis 3 1) \\<^sub>7 (axis 5 1) = axis 2 1" "(axis 5 1) \\<^sub>7 (axis 3 1) = -(axis 2 1)" "(axis 4 1) \\<^sub>7 (axis 6 1) = axis 3 1" "(axis 6 1) \\<^sub>7 (axis 4 1) = -(axis 3 1)" "(axis 5 1) \\<^sub>7 (axis 7 1) = axis 4 1" "(axis 7 1) \\<^sub>7 (axis 5 1) = -(axis 4 1)" "(axis 4 1) \\<^sub>7 (axis 1 1) = axis 2 1" "(axis 1 1) \\<^sub>7 (axis 4 1) = -(axis 2 1)" "(axis 5 1) \\<^sub>7 (axis 2 1) = axis 3 1" "(axis 2 1) \\<^sub>7 (axis 5 1) = -(axis 3 1)" "(axis 6 1) \\<^sub>7 (axis 3 1) = axis 4 1" "(axis 3 1) \\<^sub>7 (axis 6 1) = -(axis 4 1)" "(axis 7 1) \\<^sub>7 (axis 4 1) = axis 5 1" "(axis 4 1) \\<^sub>7 (axis 7 1) = -(axis 5 1)" "(axis 5 1) \\<^sub>7 (axis 6 1) = axis 1 1" "(axis 6 1) \\<^sub>7 (axis 5 1) = -(axis 1 1)" "(axis 6 1) \\<^sub>7 (axis 7 1) = axis 2 1" "(axis 7 1) \\<^sub>7 (axis 6 1) = -(axis 2 1)" "(axis 7 1) \\<^sub>7 (axis 1 1) = axis 3 1" "(axis 1 1) \\<^sub>7 (axis 7 1) = -(axis 3 1)" "(axis 6 1) \\<^sub>7 (axis 1 1) = axis 5 1" "(axis 1 1) \\<^sub>7 (axis 6 1) = -(axis 5 1)" "(axis 7 1) \\<^sub>7 (axis 2 1) = axis 6 1" "(axis 2 1) \\<^sub>7 (axis 7 1) = -(axis 6 1)" "(axis 1 1) \\<^sub>7 (axis 3 1) = axis 7 1" "(axis 3 1) \\<^sub>7 (axis 1 1) = -(axis 7 1)" "(axis 1 1) \\<^sub>7 (axis 5 1) = axis 6 1" "(axis 5 1) \\<^sub>7 (axis 1 1) = -(axis 6 1)" "(axis 2 1) \\<^sub>7 (axis 6 1) = axis 7 1" "(axis 6 1) \\<^sub>7 (axis 2 1) = -(axis 7 1)" "(axis 3 1) \\<^sub>7 (axis 7 1) = axis 1 1" "(axis 7 1) \\<^sub>7 (axis 3 1) = -(axis 1 1)" by (simp_all add: vec_eq_iff forall_7 cross7_components) lemma cross7_basis_zero: " u=0 \ (u \\<^sub>7 axis 1 1 = 0) \ (u \\<^sub>7 axis 2 1 = 0) \ (u \\<^sub>7 axis 3 1 = 0) \ (u \\<^sub>7 axis 4 1 = 0) \ (u \\<^sub>7 axis 5 1 = 0 ) \ (u \\<^sub>7 axis 6 1 = 0 ) \ (u \\<^sub>7 axis 7 1 = 0) " by auto lemma cross7_basis_nonzero: "\ (u \\<^sub>7 axis 1 1 = 0) \ \ (u \\<^sub>7 axis 2 1 = 0) \ \ (u \\<^sub>7 axis 3 1 = 0) \ \ (u \\<^sub>7 axis 4 1 = 0) \ \ (u \\<^sub>7 axis 5 1 = 0 ) \ \ (u \\<^sub>7 axis 6 1 = 0 ) \ \ (u \\<^sub>7 axis 7 1 = 0) \ u \ 0" by auto text\Pythagorean theorem/magnitude\ lemma norm_square_vec_eq: "norm x ^ 2 = (\i\UNIV. x $ i ^ 2)" by (auto simp: norm_vec_def L2_set_def intro!: sum_nonneg) lemma norm_cross7_dot_magnitude: "(norm (x \\<^sub>7 y))\<^sup>2 = (norm x)\<^sup>2 * (norm y)\<^sup>2 - (x \ y)\<^sup>2" unfolding norm_square_vec_eq sum_7 cross7_components inner_vec_def real_norm_def inner_real_def by algebra lemma cross7_eq_0: "x \\<^sub>7 y = 0 \ collinear {0, x, y}" proof - have "x \\<^sub>7 y = 0 \ norm (x \\<^sub>7 y) = 0" by simp also have "... \ (norm x * norm y)\<^sup>2 = (x \ y)\<^sup>2" using norm_cross7_dot_magnitude [of x y] by (auto simp: power_mult_distrib) also have "... \ \x \ y\ = norm x * norm y" using power2_eq_iff by (metis (mono_tags, hide_lams) abs_minus abs_norm_cancel abs_power2 norm_mult power_abs real_norm_def) also have "... \ collinear {0, x, y}" by (rule norm_cauchy_schwarz_equal) finally show ?thesis . qed lemma cross7_eq_self: "x \\<^sub>7 y = x \ x = 0" "x \\<^sub>7 y = y \ y = 0" apply (metis cross7_zero_left dot_cross7_self(1) inner_eq_zero_iff) apply (metis cross7_zero_right dot_cross7_self(2) inner_eq_zero_iff) done lemma norm_and_cross7_eq_0: "x \ y = 0 \ x \\<^sub>7 y = 0 \ x = 0 \ y = 0" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using cross7_eq_0 norm_cauchy_schwarz_equal by force next assume ?rhs then show ?lhs by auto qed lemma bilinear_cross7: "bilinear (\\<^sub>7)" apply (auto simp add: bilinear_def linear_def) apply unfold_locales apply (simp_all add: cross7_add_right cross7_mult_right cross7_add_left cross7_mult_left) done subsection\Continuity\ lemma continuous_cross7: "\continuous F f; continuous F g\ \ continuous F (\x. f x \\<^sub>7 g x)" by (subst continuous_componentwise)(auto intro!: continuous_intros simp: cross7_simps) lemma continuous_on_cross: fixes f :: "'a::t2_space \ real^7" shows "\continuous_on S f; continuous_on S g\ \ continuous_on S (\x. f x \\<^sub>7 g x)" by (simp add: continuous_on_eq_continuous_within continuous_cross7) end \ No newline at end of file diff --git a/thys/Ptolemys_Theorem/Ptolemys_Theorem.thy b/thys/Ptolemys_Theorem/Ptolemys_Theorem.thy --- a/thys/Ptolemys_Theorem/Ptolemys_Theorem.thy +++ b/thys/Ptolemys_Theorem/Ptolemys_Theorem.thy @@ -1,441 +1,441 @@ (* Author: Lukas Bulwahn *) section \Ptolemy's Theorem\ theory Ptolemys_Theorem imports - "HOL-Analysis.Ordered_Euclidean_Space" + "HOL-Analysis.Multivariate_Analysis" begin subsection \Preliminaries\ subsubsection \Additions to Rat theory\ hide_const (open) normalize subsubsection \Additions to Transcendental theory\ text \ Lemmas about @{const arcsin} and @{const arccos} commonly involve to show that their argument is in the domain of those partial functions, i.e., the argument @{term y} is between @{term "-1::real"} and @{term "1::real"}. As the argumentation for @{term "(-1::real) \ y"} and @{term "y \ (1::real)"} is often very similar, we prefer to prove @{term "\y\ \ (1::real)"} to the two goals above. The lemma for rewriting the term @{term "cos (arccos y)"} is already provided in the Isabelle distribution with name @{thm [source] cos_arccos_abs}. Here, we further provide the analogue on @{term "arcsin"} for rewriting @{term "sin (arcsin y)"}. \ lemma sin_arcsin_abs: "\y\ \ 1 \ sin (arcsin y) = y" by (simp add: abs_le_iff) text \ The further lemmas are the required variants from existing lemmas @{thm [source] arccos_lbound} and @{thm [source] arccos_ubound}. \ lemma arccos_lbound_abs [simp]: "\y\ \ 1 \ 0 \ arccos y" by (simp add: arccos_lbound) lemma arccos_ubound_abs [simp]: "\y\ \ 1 \ arccos y \ pi" by (simp add: arccos_ubound) text \ As we choose angles to be between @{term "0::real"} between @{term "2 * pi"}, we need some lemmas to reason about the sign of @{term "sin x"} for angles @{term "x"}. \ lemma sin_ge_zero_iff: assumes "0 \ x" "x < 2 * pi" shows "0 \ sin x \ x \ pi" proof assume "0 \ sin x" show "x \ pi" proof (rule ccontr) assume "\ x \ pi" from this \x < 2 * pi\ have "sin x < 0" using sin_lt_zero by auto from this \0 \ sin x\ show False by auto qed next assume "x \ pi" from this \0 \ x\ show "0 \ sin x" by (simp add: sin_ge_zero) qed lemma sin_less_zero_iff: assumes "0 \ x" "x < 2 * pi" shows "sin x < 0 \ pi < x" using assms sin_ge_zero_iff by fastforce subsubsection \Addition to Finite-Cartesian-Product theory\ text \ Here follow generally useful additions and specialised equations for two-dimensional real-valued vectors. \ lemma axis_nth_eq_0 [simp]: assumes "i \ j" shows "axis i x $ j = 0" using assms unfolding axis_def by simp lemma norm_axis: fixes x :: real shows "norm (axis i x) = abs x" by (simp add: norm_eq_sqrt_inner inner_axis_axis) lemma norm_eq_on_real_2_vec: fixes x :: "real ^ 2" shows "norm x = sqrt ((x $ 1) ^ 2 + (x $ 2) ^ 2)" by (simp add: norm_eq_sqrt_inner inner_vec_def UNIV_2 power2_eq_square) lemma dist_eq_on_real_2_vec: fixes a b :: "real ^ 2" shows "dist a b = sqrt ((a $ 1 - b $ 1) ^ 2 + (a $ 2 - b $ 2) ^ 2)" unfolding dist_norm norm_eq_on_real_2_vec by simp subsection \Polar Form of Two-Dimensional Real-Valued Vectors\ subsubsection \Definitions to Transfer to Polar Form and Back\ definition of_radiant :: "real \ real ^ 2" where "of_radiant \ = axis 1 (cos \) + axis 2 (sin \)" definition normalize :: "real ^ 2 \ real ^ 2" where "normalize p = (if p = 0 then axis 1 1 else (1 / norm p) *\<^sub>R p)" definition radiant_of :: "real ^ 2 \ real" where "radiant_of p = (THE \. 0 \ \ \ \ < 2 * pi \ of_radiant \ = normalize p)" text \ The vector @{term "of_radiant \"} is the vector with length @{term "1::real"} and angle @{term "\"} to the first axis. We normalize vectors to length @{term "1::real"} keeping their orientation with the normalize function. Conversely, @{term "radiant_of p"} is the angle of vector @{term p} to the first axis, where we choose @{term "radiant_of"} to return angles between @{term "0::real"} and @{term "2 * pi"}, following the usual high-school convention. With these definitions, we can express the main result @{term "norm p *\<^sub>R of_radiant (radiant_of p) = p"}. Note that the main result holds for any definition of @{term "radiant_of 0"}. So, we choose to define @{term "normalize 0"} and @{term "radiant_of 0"}, such that @{term "radiant_of 0 = 0"}. \ subsubsection \Lemmas on @{const of_radiant}\ lemma nth_of_radiant_1 [simp]: "of_radiant \ $ 1 = cos \" unfolding of_radiant_def by simp lemma nth_of_radiant_2 [simp]: "of_radiant \ $ 2 = sin \" unfolding of_radiant_def by simp lemma norm_of_radiant: "norm (of_radiant \) = 1" unfolding of_radiant_def norm_eq_on_real_2_vec by simp lemma of_radiant_plus_2pi: "of_radiant (\ + 2 * pi) = of_radiant \" unfolding of_radiant_def by simp lemma of_radiant_minus_2pi: "of_radiant (\ - 2 * pi) = of_radiant \" proof - have "of_radiant (\ - 2 * pi) = of_radiant (\ - 2 * pi + 2 * pi)" by (simp only: of_radiant_plus_2pi) also have "\ = of_radiant \" by simp finally show ?thesis . qed subsubsection \Lemmas on @{const normalize}\ lemma normalize_eq: "norm p *\<^sub>R normalize p = p" unfolding normalize_def by simp lemma norm_normalize: "norm (normalize p) = 1" unfolding normalize_def by (auto simp add: norm_axis) lemma nth_normalize [simp]: "\normalize p $ i\ \ 1" using norm_normalize component_le_norm_cart by metis lemma normalize_square: "(normalize p $ 1)\<^sup>2 + (normalize p $ 2)\<^sup>2 = 1" using dot_square_norm[of "normalize p"] by (simp add: inner_vec_def UNIV_2 power2_eq_square norm_normalize) lemma nth_normalize_ge_zero_iff: "0 \ normalize p $ i \ 0 \ p $ i" proof assume "0 \ normalize p $ i" from this show "0 \ p $ i" unfolding normalize_def by (auto split: if_split_asm simp add: zero_le_divide_iff) next assume "0 \ p $ i" have "0 \ axis 1 (1 :: real) $ i" using exhaust_2[of i] by auto from this \0 \ p $ i\ show "0 \ normalize p $ i" unfolding normalize_def by auto qed lemma nth_normalize_less_zero_iff: "normalize p $ i < 0 \ p $ i < 0" using nth_normalize_ge_zero_iff leD leI by metis lemma normalize_boundary_iff: "\normalize p $ 1\ = 1 \ p $ 2 = 0" proof assume "\normalize p $ 1\ = 1" from this have 1: "(p $ 1) ^ 2 = norm p ^ 2" unfolding normalize_def by (auto split: if_split_asm simp add: power2_eq_iff) moreover have "(p $ 1) ^ 2 + (p $ 2) ^ 2 = norm p ^ 2" using norm_eq_on_real_2_vec by auto ultimately show "p $ 2 = 0" by simp next assume "p $ 2 = 0" from this have "\p $ 1\ = norm p" by (auto simp add: norm_eq_on_real_2_vec) from this show "\normalize p $ 1\ = 1" unfolding normalize_def by simp qed lemma between_normalize_if_distant_from_0: assumes "norm p \ 1" shows "between (0, p) (normalize p)" using assms by (auto simp add: between_mem_segment closed_segment_def normalize_def) lemma between_normalize_if_near_0: assumes "norm p \ 1" shows "between (0, normalize p) p" proof - have "0 \ norm p" by simp from assms have "p = (norm p / norm p) *\<^sub>R p \ 0 \ norm p \ norm p \ 1" by auto from this have "\u. p = (u / norm p) *\<^sub>R p \ 0 \ u \ u \ 1" by blast from this show ?thesis by (auto simp add: between_mem_segment closed_segment_def normalize_def) qed subsubsection \Lemmas on @{const radiant_of}\ lemma radiant_of: "0 \ radiant_of p \ radiant_of p < 2 * pi \ of_radiant (radiant_of p) = normalize p" proof - let ?a = "if 0 \ p $ 2 then arccos (normalize p $ 1) else pi + arccos (- (normalize p $ 1))" have "0 \ ?a \ ?a < 2 * pi \ of_radiant ?a = normalize p" proof - have "0 \ ?a" by auto moreover have "?a < 2 * pi" proof cases assume "0 \ p $ 2" from this have "?a \ pi" by simp from this show ?thesis using pi_gt_zero by linarith next assume "\ 0 \ p $ 2" have "arccos (- normalize p $ 1) < pi" proof - have "\normalize p $ 1\ \ 1" using \\ 0 \ p $ 2\ by (simp only: normalize_boundary_iff) from this have "arccos (- normalize p $ 1) \ pi" unfolding arccos_minus_1[symmetric] by (subst arccos_eq_iff) auto moreover have "arccos (- normalize p $ 1) \ pi" by simp ultimately show "arccos (- normalize p $ 1) < pi" by linarith qed from this \\ 0 \ p $ 2\ show ?thesis by simp qed moreover have "of_radiant ?a = normalize p" proof - have "of_radiant ?a $ i = normalize p $ i" for i proof - have "of_radiant ?a $ 1 = normalize p $ 1" unfolding of_radiant_def by (simp add: cos_arccos_abs) moreover have "of_radiant ?a $ 2 = normalize p $ 2" proof cases assume "0 \ p $ 2" have "sin (arccos (normalize p $ 1)) = sqrt (1 - (normalize p $ 1) ^ 2)" by (simp add: sin_arccos_abs) also have "\ = normalize p $ 2" proof - have "1 - (normalize p $ 1)\<^sup>2 = (normalize p $ 2)\<^sup>2" using normalize_square[of p] by auto from this \0 \ p $ 2\ show ?thesis by (simp add: nth_normalize_ge_zero_iff) qed finally show ?thesis using \0 \ p $ 2\ unfolding of_radiant_def by auto next assume "\ 0 \ p $ 2" have "- sin (arccos (- normalize p $ 1)) = - sqrt (1 - (normalize p $ 1)\<^sup>2)" by (simp add: sin_arccos_abs) also have "\ = normalize p $ 2" proof - have "1 - (normalize p $ 1)\<^sup>2 = (normalize p $ 2)\<^sup>2" using normalize_square[of p] by auto from this \\ 0 \ p $ 2\ show ?thesis using nth_normalize_ge_zero_iff by fastforce qed finally show ?thesis using \\ 0 \ p $ 2\ unfolding of_radiant_def by auto qed ultimately show ?thesis by (metis exhaust_2[of i]) qed from this show ?thesis by (simp add: vec_eq_iff) qed ultimately show ?thesis by blast qed moreover { fix \ assume "0 \ \ \ \ < 2 * pi \ of_radiant \ = normalize p" from this have "0 \ \" "\ < 2 * pi" "normalize p = of_radiant \" by auto from this have "cos \ = normalize p $ 1" "sin \ = normalize p $ 2" by auto have "\ = ?a" proof cases assume "0 \ p $ 2" from this have "\ \ pi" using \0 \ \\ \\ < 2 * pi\ \sin \ = normalize p $ 2\ by (simp add: sin_ge_zero_iff[symmetric] nth_normalize_ge_zero_iff) from \0 \ \\ this have "\ = arccos (cos \)" by (simp add: arccos_cos) from \cos \ = normalize p $ 1\ this have "\ = arccos (normalize p $ 1)" by (simp add: arccos_eq_iff) from this show "\ = ?a" using \0 \ p $ 2\ by auto next assume "\ 0 \ p $ 2" from this have "\ > pi" using \0 \ \\ \\ < 2 * pi\ \sin \ = normalize p $ 2\ by (simp add: sin_less_zero_iff[symmetric] nth_normalize_less_zero_iff) from this \\ < 2 * pi\ have "\ - pi = arccos (cos (\ - pi))" by (auto simp only: arccos_cos) from this \cos \ = normalize p $ 1\ have "\ - pi = arccos (- normalize p $ 1)" by simp from this have "\ = pi + arccos (- normalize p $ 1)" by simp from this show "\ = ?a" using \\ 0 \ p $ 2\ by auto qed } ultimately show ?thesis unfolding radiant_of_def by (rule theI) qed lemma radiant_of_bounds [simp]: "0 \ radiant_of p" "radiant_of p < 2 * pi" using radiant_of by auto lemma radiant_of_weak_ubound [simp]: "radiant_of p \ 2 * pi" using radiant_of_bounds(2)[of p] by linarith subsubsection \Main Equations for Transforming to Polar Form\ lemma polar_form_eq: "norm p *\<^sub>R of_radiant (radiant_of p) = p" using radiant_of normalize_eq by simp lemma relative_polar_form_eq: "Q + dist P Q *\<^sub>R of_radiant (radiant_of (P - Q)) = P" proof - have "norm (P - Q) *\<^sub>R of_radiant (radiant_of (P - Q)) = P - Q" unfolding polar_form_eq .. moreover have "dist P Q = norm (P - Q)" by (simp add: dist_norm) ultimately show ?thesis by (metis add.commute diff_add_cancel) qed subsection \Ptolemy's Theorem\ lemma dist_circle_segment: assumes "0 \ radius" "0 \ \" "\ \ \" "\ \ 2 * pi" shows "dist (center + radius *\<^sub>R of_radiant \) (center + radius *\<^sub>R of_radiant \) = 2 * radius * sin ((\ - \) / 2)" (is "?lhs = ?rhs") proof - have trigonometry: "(cos \ - cos \)\<^sup>2 + (sin \ - sin \)\<^sup>2 = (2 * sin ((\ - \) / 2))\<^sup>2" proof - have sin_diff_minus: "sin ((\ - \) / 2) = - sin ((\ - \) / 2)" by (simp only: sin_minus[symmetric] minus_divide_left minus_diff_eq) have "(cos \ - cos \)\<^sup>2 + (sin \ - sin \)\<^sup>2 = (2 * sin ((\ + \) / 2) * sin ((\ - \) / 2))\<^sup>2 + (2 * sin ((\ - \) / 2) * cos ((\ + \) / 2))\<^sup>2" by (simp only: cos_diff_cos sin_diff_sin) also have "\ = (2 * sin ((\ - \) / 2))\<^sup>2 * ((sin ((\ + \) / 2))\<^sup>2 + (cos ((\ + \) / 2))\<^sup>2)" unfolding sin_diff_minus by algebra also have "\ = (2 * sin ((\ - \) / 2))\<^sup>2" by simp finally show ?thesis . qed from assms have "0 \ sin ((\ - \) / 2)" by (simp add: sin_ge_zero) have "?lhs = sqrt (radius\<^sup>2 * ((cos \ - cos \)\<^sup>2 + (sin \ - sin \)\<^sup>2))" unfolding dist_eq_on_real_2_vec by simp algebra also have "\ = sqrt (radius\<^sup>2 * (2 * sin ((\ - \) / 2))\<^sup>2)" by (simp add: trigonometry) also have "\ = ?rhs" using \0 \ radius\ \0 \ sin ((\ - \) / 2)\ by (simp add: real_sqrt_mult) finally show ?thesis . qed theorem ptolemy_trigonometric: fixes \\<^sub>1 \\<^sub>2 \\<^sub>3 :: real shows "sin (\\<^sub>1 + \\<^sub>2) * sin (\\<^sub>2 + \\<^sub>3) = sin \\<^sub>1 * sin \\<^sub>3 + sin \\<^sub>2 * sin (\\<^sub>1 + \\<^sub>2 + \\<^sub>3)" proof - have "sin (\\<^sub>1 + \\<^sub>2) * sin (\\<^sub>2 + \\<^sub>3) = ((sin \\<^sub>2)\<^sup>2 + (cos \\<^sub>2)\<^sup>2) * sin \\<^sub>1 * sin \\<^sub>3 + sin \\<^sub>2 * sin (\\<^sub>1 + \\<^sub>2 + \\<^sub>3)" by (simp only: sin_add cos_add) algebra also have "\ = sin \\<^sub>1 * sin \\<^sub>3 + sin \\<^sub>2 * sin (\\<^sub>1 + \\<^sub>2 + \\<^sub>3)" by simp finally show ?thesis . qed theorem ptolemy: fixes A B C D center :: "real ^ 2" assumes "dist center A = radius" and "dist center B = radius" assumes "dist center C = radius" and "dist center D = radius" assumes ordering_of_points: "radiant_of (A - center) \ radiant_of (B - center)" "radiant_of (B - center) \ radiant_of (C - center)" "radiant_of (C - center) \ radiant_of (D - center)" shows "dist A C * dist B D = dist A B * dist C D + dist A D * dist B C" proof - from \dist center A = radius\ have "0 \ radius" by auto define \ \ \ \ where "\ = radiant_of (A - center)" and "\ = radiant_of (B - center)" and "\ = radiant_of (C - center)" and "\ = radiant_of (D - center)" from ordering_of_points have angle_basics: "\ \ \" "\ \ \" "\ \ \" "0 \ \" "\ \ 2 * pi" "0 \ \" "\ \ 2 * pi" "0 \ \" "\ \ 2 * pi" "0 \ \" "\ \ 2 * pi" unfolding \_def \_def \_def \_def by auto from assms(1-4) have "A = center + radius *\<^sub>R of_radiant \" "B = center + radius *\<^sub>R of_radiant \" "C = center + radius *\<^sub>R of_radiant \" "D = center + radius *\<^sub>R of_radiant \" unfolding \_def \_def \_def \_def using relative_polar_form_eq dist_commute by metis+ from this have dist_eqs: "dist A C = 2 * radius * sin ((\ - \) / 2)" "dist B D = 2 * radius * sin ((\ - \) / 2)" "dist A B = 2 * radius * sin ((\ - \) / 2)" "dist C D = 2 * radius * sin ((\ - \) / 2)" "dist A D = 2 * radius * sin ((\ - \) / 2)" "dist B C = 2 * radius * sin ((\ - \) / 2)" using angle_basics \radius \ 0\ dist_circle_segment by (auto) have "dist A C * dist B D = 4 * radius ^ 2 * sin ((\ - \) / 2) * sin ((\ - \) / 2)" unfolding dist_eqs by (simp add: power2_eq_square) also have "\ = 4 * radius ^ 2 * (sin ((\ - \) / 2) * sin ((\ - \) / 2) + sin ((\ - \) / 2) * sin ((\ - \) / 2))" proof - define \\<^sub>1 \\<^sub>2 \\<^sub>3 where "\\<^sub>1 = (\ - \) / 2" and "\\<^sub>2 = (\ - \) / 2" and "\\<^sub>3 = (\ - \) / 2" have "(\ - \) / 2 = \\<^sub>1 + \\<^sub>2" and "(\ - \) / 2 = \\<^sub>2 + \\<^sub>3" and "(\ - \) / 2 = \\<^sub>1 + \\<^sub>2 + \\<^sub>3" unfolding \\<^sub>1_def \\<^sub>2_def \\<^sub>3_def by (auto simp add: field_simps) have "sin ((\ - \) / 2) * sin ((\ - \) / 2) = sin (\\<^sub>1 + \\<^sub>2) * sin (\\<^sub>2 + \\<^sub>3)" using \(\ - \) / 2 = \\<^sub>1 + \\<^sub>2\ \(\ - \) / 2 = \\<^sub>2 + \\<^sub>3\ by (simp only:) also have "\ = sin \\<^sub>1 * sin \\<^sub>3 + sin \\<^sub>2 * sin (\\<^sub>1 + \\<^sub>2 + \\<^sub>3)" by (rule ptolemy_trigonometric) also have "\ = (sin ((\ - \) / 2) * sin ((\ - \) / 2) + sin ((\ - \) / 2) * sin ((\ - \) / 2))" using \\<^sub>1_def \\<^sub>2_def \\<^sub>3_def \(\ - \) / 2 = \\<^sub>1 + \\<^sub>2 + \\<^sub>3\ by (simp only:) finally show ?thesis by simp qed also have "\ = dist A B * dist C D + dist A D * dist B C" unfolding dist_eqs by (simp add: distrib_left power2_eq_square) finally show ?thesis . qed end diff --git a/thys/Quaternions/Quaternions.thy b/thys/Quaternions/Quaternions.thy --- a/thys/Quaternions/Quaternions.thy +++ b/thys/Quaternions/Quaternions.thy @@ -1,748 +1,747 @@ section\Theory of Quaternions\ text\This theory is inspired by the HOL Light development of quaternions, but follows its own route. Quaternions are developed coinductively, as in the existing formalisation of the complex numbers. Quaternions are quickly shown to belong to the type classes of real normed division algebras and real inner product spaces. And therefore they inherit a great body of facts involving algebraic laws, limits, continuity, etc., which must be proved explicitly in the HOL Light version. The development concludes with the geometric interpretation of the product of imaginary quaternions.\ theory "Quaternions" imports - "HOL-Analysis.Ordered_Euclidean_Space" - "HOL-Analysis.Cross3" + "HOL-Analysis.Multivariate_Analysis" begin subsection\Basic definitions\ text\As with the complex numbers, coinduction is convenient\ codatatype quat = Quat (Re: real) (Im1: real) (Im2: real) (Im3: real) lemma quat_eqI [intro?]: "\Re x = Re y; Im1 x = Im1 y; Im2 x = Im2 y; Im3 x = Im3 y\ \ x = y" by (rule quat.expand) simp lemma quat_eq_iff: "x = y \ Re x = Re y \ Im1 x = Im1 y \ Im2 x = Im2 y \ Im3 x = Im3 y" by (auto intro: quat.expand) context begin no_notation Complex.imaginary_unit ("\") primcorec quat_ii :: quat ("\") where "Re \ = 0" | "Im1 \ = 1" | "Im2 \ = 0" | "Im3 \ = 0" primcorec quat_jj :: quat ("\") where "Re \ = 0" | "Im1 \ = 0" | "Im2 \ = 1" | "Im3 \ = 0" primcorec quat_kk :: quat ("\") where "Re \ = 0" | "Im1 \ = 0" | "Im2 \ = 0" | "Im3 \ = 1" end bundle quaternion_syntax begin notation quat_ii ("\") no_notation Complex.imaginary_unit ("\") end bundle no_quaternion_syntax begin no_notation quat_ii ("\") notation Complex.imaginary_unit ("\") end unbundle quaternion_syntax subsection \Addition and Subtraction: An Abelian Group\ instantiation quat :: ab_group_add begin primcorec zero_quat where "Re 0 = 0" |"Im1 0 = 0" | "Im2 0 = 0" | "Im3 0 = 0" primcorec plus_quat where "Re (x + y) = Re x + Re y" | "Im1 (x + y) = Im1 x + Im1 y" | "Im2 (x + y) = Im2 x + Im2 y" | "Im3 (x + y) = Im3 x + Im3 y" primcorec uminus_quat where "Re (- x) = - Re x" | "Im1 (- x) = - Im1 x" | "Im2 (- x) = - Im2 x" | "Im3 (- x) = - Im3 x" primcorec minus_quat where "Re (x - y) = Re x - Re y" | "Im1 (x - y) = Im1 x - Im1 y" | "Im2 (x - y) = Im2 x - Im2 y" | "Im3 (x - y) = Im3 x - Im3 y" instance by standard (simp_all add: quat_eq_iff) end subsection \A Vector Space\ instantiation quat :: real_vector begin primcorec scaleR_quat where "Re (scaleR r x) = r * Re x" | "Im1 (scaleR r x) = r * Im1 x" | "Im2 (scaleR r x) = r * Im2 x" | "Im3 (scaleR r x) = r * Im3 x" instance by standard (auto simp: quat_eq_iff distrib_left distrib_right scaleR_add_right) end instantiation quat :: real_algebra_1 begin primcorec one_quat where "Re 1 = 1" | "Im1 1 = 0" | "Im2 1 = 0" | "Im3 1 = 0" primcorec times_quat where "Re (x * y) = Re x * Re y - Im1 x * Im1 y - Im2 x * Im2 y - Im3 x * Im3 y" | "Im1 (x * y) = Re x * Im1 y + Im1 x * Re y + Im2 x * Im3 y - Im3 x * Im2 y" | "Im2 (x * y) = Re x * Im2 y - Im1 x * Im3 y + Im2 x * Re y + Im3 x * Im1 y" | "Im3 (x * y) = Re x * Im3 y + Im1 x * Im2 y - Im2 x * Im1 y + Im3 x * Re y" instance by standard (auto simp: quat_eq_iff distrib_left distrib_right Rings.right_diff_distrib Rings.left_diff_distrib) end subsection \Multiplication and Division: A Real Division Algebra\ instantiation quat :: real_div_algebra begin primcorec inverse_quat where "Re (inverse x) = Re x / ((Re x)\<^sup>2 + (Im1 x)\<^sup>2 + (Im2 x)\<^sup>2 + (Im3 x)\<^sup>2)" | "Im1 (inverse x) = - (Im1 x) / ((Re x)\<^sup>2 + (Im1 x)\<^sup>2 + (Im2 x)\<^sup>2 + (Im3 x)\<^sup>2)" | "Im2 (inverse x) = - (Im2 x) / ((Re x)\<^sup>2 + (Im1 x)\<^sup>2 + (Im2 x)\<^sup>2 + (Im3 x)\<^sup>2)" | "Im3 (inverse x) = - (Im3 x) / ((Re x)\<^sup>2 + (Im1 x)\<^sup>2 + (Im2 x)\<^sup>2 + (Im3 x)\<^sup>2)" definition "x div y = x * inverse y" for x y :: quat instance proof show "\x::quat. x \ 0 \ inverse x * x = 1" by (auto simp: quat_eq_iff add_nonneg_eq_0_iff power2_eq_square add_divide_distrib [symmetric] diff_divide_distrib [symmetric]) show "\x::quat. x \ 0 \ x * inverse x = 1" by (auto simp: quat_eq_iff add_nonneg_eq_0_iff power2_eq_square add_divide_distrib [symmetric]) show "\x y::quat. x div y = x * inverse y" by (simp add: divide_quat_def) show "inverse 0 = (0::quat)" by (auto simp: quat_eq_iff) qed end subsection \Multiplication and Division: A Real Normed Division Algebra\ fun quat_proj where "quat_proj x 0 = Re x" | "quat_proj x (Suc 0) = Im1 x" | "quat_proj x (Suc (Suc 0)) = Im2 x" | "quat_proj x (Suc (Suc (Suc 0))) = Im3 x" lemma quat_proj_add: assumes "i \ 3" shows "quat_proj (x+y) i = quat_proj x i + quat_proj y i" proof - consider "i = 0" | "i = 1" | "i = 2" | "i = 3" using assms by linarith then show ?thesis by cases (auto simp: numeral_2_eq_2 numeral_3_eq_3) qed instantiation quat :: real_normed_div_algebra begin definition "norm z = sqrt ((Re z)\<^sup>2 + (Im1 z)\<^sup>2 + (Im2 z)\<^sup>2 + (Im3 z)\<^sup>2)" definition "sgn x = x /\<^sub>R norm x" for x :: quat definition "dist x y = norm (x - y)" for x y :: quat definition [code del]: "(uniformity :: (quat \ quat) filter) = (INF e\{0 <..}. principal {(x, y). dist x y < e})" definition [code del]: "open (U :: quat set) \ (\x\U. eventually (\(x', y). x' = x \ y \ U) uniformity)" lemma norm_eq_L2: "norm z = L2_set (quat_proj z) {..3}" by (simp add: norm_quat_def L2_set_def numeral_3_eq_3) instance proof fix r :: real and x y :: quat and S :: "quat set" show "(norm x = 0) \ (x = 0)" by (simp add: norm_quat_def quat_eq_iff add_nonneg_eq_0_iff) have eq: "L2_set (quat_proj (x + y)) {..3} = L2_set (\i. quat_proj x i + quat_proj y i) {..3}" by (rule L2_set_cong) (auto simp: quat_proj_add) show "norm (x + y) \ norm x + norm y" by (simp add: norm_eq_L2 eq L2_set_triangle_ineq) show "norm (scaleR r x) = \r\ * norm x" by (simp add: norm_quat_def quat_eq_iff power_mult_distrib distrib_left [symmetric] real_sqrt_mult) show "norm (x * y) = norm x * norm y" by (simp add: norm_quat_def quat_eq_iff real_sqrt_mult [symmetric] power2_eq_square algebra_simps) qed (rule sgn_quat_def dist_quat_def open_quat_def uniformity_quat_def)+ end instantiation quat :: real_inner begin definition inner_quat_def: "inner x y = Re x * Re y + Im1 x * Im1 y + Im2 x * Im2 y + Im3 x * Im3 y" instance proof fix x y z :: quat and r :: real show "inner x y = inner y x" unfolding inner_quat_def by (simp add: mult.commute) show "inner (x + y) z = inner x z + inner y z" unfolding inner_quat_def by (simp add: distrib_right) show "inner (scaleR r x) y = r * inner x y" unfolding inner_quat_def by (simp add: distrib_left) show "0 \ inner x x" unfolding inner_quat_def by simp show "inner x x = 0 \ x = 0" unfolding inner_quat_def by (simp add: add_nonneg_eq_0_iff quat_eq_iff) show "norm x = sqrt (inner x x)" unfolding inner_quat_def norm_quat_def by (simp add: power2_eq_square) qed end lemma quat_inner_1 [simp]: "inner 1 x = Re x" unfolding inner_quat_def by simp lemma quat_inner_1_right [simp]: "inner x 1 = Re x" unfolding inner_quat_def by simp lemma quat_inner_i_left [simp]: "inner \ x = Im1 x" unfolding inner_quat_def by simp lemma quat_inner_i_right [simp]: "inner x \ = Im1 x" unfolding inner_quat_def by simp lemma quat_inner_j_left [simp]: "inner \ x = Im2 x" unfolding inner_quat_def by simp lemma quat_inner_j_right [simp]: "inner x \ = Im2 x" unfolding inner_quat_def by simp lemma quat_inner_k_left [simp]: "inner \ x = Im3 x" unfolding inner_quat_def by simp lemma quat_inner_k_right [simp]: "inner x \ = Im3 x" unfolding inner_quat_def by simp abbreviation quat_of_real :: "real \ quat" where "quat_of_real \ of_real" lemma Re_quat_of_real [simp]: "Re(quat_of_real a) = a" by (simp add: of_real_def) lemma Im1_quat_of_real [simp]: "Im1(quat_of_real a) = 0" by (simp add: of_real_def) lemma Im2_quat_of_real [simp]: "Im2(quat_of_real a) = 0" by (simp add: of_real_def) lemma Im3_quat_of_real [simp]: "Im3(quat_of_real a) = 0" by (simp add: of_real_def) lemma quat_eq_0_iff: "q = 0 \ (Re q)\<^sup>2 + (Im1 q)\<^sup>2 + (Im2 q)\<^sup>2 + (Im3 q)\<^sup>2 = 0" proof assume "(quat.Re q)\<^sup>2 + (Im1 q)\<^sup>2 + (Im2 q)\<^sup>2 + (Im3 q)\<^sup>2 = 0" then have "\qa. qa - q = qa" by (simp add: add_nonneg_eq_0_iff minus_quat.ctr) then show "q = 0" by simp qed auto lemma quat_of_real_times_commute: "quat_of_real r * q = q * of_real r" by (simp add: of_real_def) lemma quat_of_real_times_left_commute: "quat_of_real r * (p * q) = p * (of_real r * q)" by (simp add: of_real_def) lemma quat_norm_units [simp]: "norm quat_ii = 1" "norm (\::quat) = 1" "norm (\::quat) = 1" by (auto simp: norm_quat_def) lemma ii_nz [simp]: "quat_ii \ 0" using quat_ii.simps(2) by fastforce lemma jj_nz [simp]: "\ \ 0" using quat_jj.sel(3) by fastforce lemma kk_nz [simp]: "\ \ 0" using quat_kk.sel(4) by force text\An "expansion" theorem into the traditional notation\ lemma quat_unfold: "q = of_real(Re q) + \ * of_real(Im1 q) + \ * of_real(Im2 q) + \ * of_real(Im3 q)" by (simp add: quat_eq_iff) lemma quat_trad: "Quat x y z w = of_real x + \ * of_real y + \ * of_real z + \ * of_real w" by (simp add: quat_eq_iff) lemma of_real_eq_Quat: "of_real a = Quat a 0 0 0" by (simp add: quat_trad) lemma ii_squared [simp]: "quat_ii\<^sup>2 = -1" by (simp add: power2_eq_square quat.expand) lemma jj_squared [simp]: "\^2 = -1" by (simp add: power2_eq_square quat.expand) lemma kk_squared [simp]: "\^2 = -1" by (simp add: power2_eq_square quat.expand) lemma inverse_ii [simp]: "inverse quat_ii = -quat_ii" by (simp add: power2_eq_square quat.expand) lemma inverse_jj [simp]: "inverse \ = -\" by (simp add: power2_eq_square quat.expand) lemma inverse_kk [simp]: "inverse \ = -\" by (simp add: power2_eq_square quat.expand) lemma inverse_mult: "inverse (p * q) = inverse q * inverse p" for p::quat by (metis inverse_zero mult_not_zero nonzero_inverse_mult_distrib) lemma quat_of_real_inverse_collapse [simp]: assumes "c \ 0" shows "quat_of_real c * quat_of_real (inverse c) = 1" "quat_of_real (inverse c) * quat_of_real c = 1" using assms by auto subsection\Conjugate of a quaternion\ primcorec cnj :: "quat \ quat" where "Re (cnj z) = Re z" | "Im1 (cnj z) = - Im1 z" | "Im2 (cnj z) = - Im2 z" | "Im3 (cnj z) = - Im3 z" lemma cnj_cancel_iff [simp]: "cnj x = cnj y \ x = y" proof show "cnj x = cnj y \ x = y" by (simp add: quat_eq_iff) qed auto lemma cnj_cnj [simp]: "cnj(cnj q) = q" by (simp add: quat_eq_iff) lemma cnj_of_real [simp]: "cnj(quat_of_real x) = quat_of_real x" by (simp add: quat_eqI) lemma cnj_zero [simp]: "cnj 0 = 0" by (simp add: quat_eq_iff) lemma cnj_zero_iff [iff]: "cnj z = 0 \ z = 0" by (simp add: quat_eq_iff) lemma cnj_one [simp]: "cnj 1 = 1" by (simp add: quat_eq_iff) lemma cnj_one_iff [simp]: "cnj z = 1 \ z = 1" by (simp add: quat_eq_iff) lemma quat_norm_cnj [simp]: "norm(cnj q) = norm q" by (simp add: norm_quat_def) lemma cnj_add [simp]: "cnj (x + y) = cnj x + cnj y" by (simp add: quat_eq_iff) lemma cnj_sum [simp]: "cnj (sum f S) = (\x\S. cnj (f x))" by (induct S rule: infinite_finite_induct) auto lemma cnj_diff [simp]: "cnj (x - y) = cnj x - cnj y" by (simp add: quat_eq_iff) lemma cnj_minus [simp]: "cnj (- x) = - cnj x" by (simp add: quat_eq_iff) lemma cnj_mult [simp]: "cnj (x * y) = cnj y * cnj x" by (simp add: quat_eq_iff) lemma cnj_inverse [simp]: "cnj (inverse x) = inverse (cnj x)" by (simp add: quat_eq_iff) lemma cnj_divide [simp]: "cnj (x / y) = inverse (cnj y) * cnj x" by (simp add: divide_quat_def) lemma cnj_power [simp]: "cnj (x^n) = cnj x ^ n" by (induct n) (auto simp: power_commutes) lemma cnj_of_nat [simp]: "cnj (of_nat n) = of_nat n" by (metis cnj_of_real of_real_of_nat_eq) lemma cnj_of_int [simp]: "cnj (of_int z) = of_int z" by (metis cnj_of_real of_real_of_int_eq) lemma cnj_numeral [simp]: "cnj (numeral w) = numeral w" by (metis of_nat_numeral cnj_of_nat) lemma cnj_neg_numeral [simp]: "cnj (- numeral w) = - numeral w" by simp lemma cnj_scaleR [simp]: "cnj (scaleR r x) = scaleR r (cnj x)" by (simp add: quat_eq_iff) lemma cnj_units [simp]: "cnj quat_ii = -\" "cnj \ = -\" "cnj \ = -\" by (simp_all add: quat_eq_iff) lemma cnj_eq_of_real: "cnj q = quat_of_real x \ q = quat_of_real x" proof show "cnj q = quat_of_real x \ q = quat_of_real x" by (metis cnj_of_real cnj_cnj) qed auto lemma quat_add_cnj: "q + cnj q = quat_of_real(2 * Re q)" "cnj q + q = quat_of_real(2 * Re q)" by simp_all (simp_all add: mult.commute mult_2 plus_quat.code) lemma quat_divide_numeral: fixes x::quat shows "x / numeral w = x /\<^sub>R numeral w" unfolding divide_quat_def by (metis mult.right_neutral mult_scaleR_right of_real_def of_real_inverse of_real_numeral) lemma Re_divide_numeral [simp]: "Re (x / numeral w) = Re x / numeral w" by (metis divide_inverse_commute quat_divide_numeral scaleR_quat.simps(1)) lemma Im1_divide_numeral [simp]: "Im1 (x / numeral w) = Im1 x / numeral w" unfolding quat_divide_numeral by simp lemma Im2_divide_numeral [simp]: "Im2 (x / numeral w) = Im2 x / numeral w" unfolding quat_divide_numeral by simp lemma Im3_divide_numeral [simp]: "Im3 (x / numeral w) = Im3 x / numeral w" unfolding quat_divide_numeral by simp lemma of_real_quat_re_cnj: "quat_of_real(Re q) = inverse(quat_of_real 2) * (q + cnj q)" by (simp add: quat_eq_iff) lemma quat_mult_cnj_commute: "cnj q * q = q * cnj q" by (simp add: quat_eq_iff) lemma quat_norm_pow_2: "quat_of_real(norm q) ^ 2 = q * cnj q" by (simp add: quat_eq_iff norm_quat_def flip: of_real_power) (auto simp: power2_eq_square) lemma quat_norm_pow_2_alt: "quat_of_real(norm q) ^ 2 = cnj q * q" by (simp add: quat_mult_cnj_commute quat_norm_pow_2) lemma quat_inverse_cnj: "inverse q = inverse (quat_of_real((norm q)\<^sup>2)) * cnj q" by (simp add: quat_eq_iff norm_quat_def numeral_Bit0 flip: of_real_power) lemma quat_inverse_eq_cnj: "norm q = 1 \ inverse q = cnj q" by (metis inverse_1 mult_cancel_left norm_eq_zero norm_one cnj_one quat_norm_pow_2 right_inverse) subsection\Linearity and continuity of the components\ lemma bounded_linear_Re: "bounded_linear Re" and bounded_linear_Im1: "bounded_linear Im1" and bounded_linear_Im2: "bounded_linear Im2" and bounded_linear_Im3: "bounded_linear Im3" by (simp_all add: bounded_linear_intro [where K=1] norm_quat_def real_le_rsqrt add.assoc) lemmas Cauchy_Re = bounded_linear.Cauchy [OF bounded_linear_Re] lemmas Cauchy_Im1 = bounded_linear.Cauchy [OF bounded_linear_Im1] lemmas Cauchy_Im2 = bounded_linear.Cauchy [OF bounded_linear_Im2] lemmas Cauchy_Im3 = bounded_linear.Cauchy [OF bounded_linear_Im3] lemmas tendsto_Re [tendsto_intros] = bounded_linear.tendsto [OF bounded_linear_Re] lemmas tendsto_Im1 [tendsto_intros] = bounded_linear.tendsto [OF bounded_linear_Im1] lemmas tendsto_Im2 [tendsto_intros] = bounded_linear.tendsto [OF bounded_linear_Im2] lemmas tendsto_Im3 [tendsto_intros] = bounded_linear.tendsto [OF bounded_linear_Im3] lemmas isCont_Re [simp] = bounded_linear.isCont [OF bounded_linear_Re] lemmas isCont_Im1 [simp] = bounded_linear.isCont [OF bounded_linear_Im1] lemmas isCont_Im2 [simp] = bounded_linear.isCont [OF bounded_linear_Im2] lemmas isCont_Im3 [simp] = bounded_linear.isCont [OF bounded_linear_Im3] lemmas continuous_Re [simp] = bounded_linear.continuous [OF bounded_linear_Re] lemmas continuous_Im1 [simp] = bounded_linear.continuous [OF bounded_linear_Im1] lemmas continuous_Im2 [simp] = bounded_linear.continuous [OF bounded_linear_Im2] lemmas continuous_Im3 [simp] = bounded_linear.continuous [OF bounded_linear_Im3] lemmas continuous_on_Re [continuous_intros] = bounded_linear.continuous_on[OF bounded_linear_Re] lemmas continuous_on_Im1 [continuous_intros] = bounded_linear.continuous_on[OF bounded_linear_Im1] lemmas continuous_on_Im2 [continuous_intros] = bounded_linear.continuous_on[OF bounded_linear_Im2] lemmas continuous_on_Im3 [continuous_intros] = bounded_linear.continuous_on[OF bounded_linear_Im3] lemmas has_derivative_Re [derivative_intros] = bounded_linear.has_derivative[OF bounded_linear_Re] lemmas has_derivative_Im1 [derivative_intros] = bounded_linear.has_derivative[OF bounded_linear_Im1] lemmas has_derivative_Im2 [derivative_intros] = bounded_linear.has_derivative[OF bounded_linear_Im2] lemmas has_derivative_Im3 [derivative_intros] = bounded_linear.has_derivative[OF bounded_linear_Im3] lemmas sums_Re = bounded_linear.sums [OF bounded_linear_Re] lemmas sums_Im1 = bounded_linear.sums [OF bounded_linear_Im1] lemmas sums_Im2 = bounded_linear.sums [OF bounded_linear_Im2] lemmas sums_Im3 = bounded_linear.sums [OF bounded_linear_Im3] subsection\Quaternionic-specific theorems about sums\ lemma Re_sum [simp]: "Re(sum f S) = sum (\x. Re(f x)) S" for f :: "'a \ quat" by (induct S rule: infinite_finite_induct) auto lemma Im1_sum [simp]: "Im1(sum f S) = sum (\x. Im1(f x)) S" by (induct S rule: infinite_finite_induct) auto lemma Im2_sum [simp]: "Im2(sum f S) = sum (\x. Im2(f x)) S" by (induct S rule: infinite_finite_induct) auto lemma Im3_sum [simp]: "Im3(sum f S) = sum (\x. Im3(f x)) S" by (induct S rule: infinite_finite_induct) auto lemma in_Reals_iff_Re: "q \ Reals \ quat_of_real(Re q) = q" by (metis Re_quat_of_real Reals_cases Reals_of_real) lemma in_Reals_iff_cnj: "q \ Reals \ cnj q = q" by (simp add: in_Reals_iff_Re quat_eq_iff) lemma real_norm: "q \ Reals \ norm q = abs(Re q)" by (metis in_Reals_iff_Re norm_of_real) lemma norm_power2: "(norm q)\<^sup>2 = Re(cnj q * q)" by (metis Re_quat_of_real of_real_power quat_mult_cnj_commute quat_norm_pow_2) lemma quat_norm_imaginary: "Re x = 0 \ x\<^sup>2 = - (quat_of_real (norm x))\<^sup>2" unfolding quat_norm_pow_2 by (cases x) (auto simp: power2_eq_square cnj.ctr times_quat.ctr uminus_quat.ctr) subsection\Bound results for real and imaginary components of limits\ lemma Re_tendsto_upperbound: "\(f \ l) net; \\<^sub>F x in net. quat.Re (f x) \ b; net \ bot\ \ Re l \ b" by (blast intro: tendsto_upperbound [OF tendsto_Re]) lemma Im1_tendsto_upperbound: "\(f \ l) net; \\<^sub>F x in net. Im1 (f x) \ b; net \ bot\ \ Im1 l \ b" by (blast intro: tendsto_upperbound [OF tendsto_Im1]) lemma Im2_tendsto_upperbound: "\(f \ l) net; \\<^sub>F x in net. Im2 (f x) \ b; net \ bot\ \ Im2 l \ b" by (blast intro: tendsto_upperbound [OF tendsto_Im2]) lemma Im3_tendsto_upperbound: "\(f \ l) net; \\<^sub>F x in net. Im3 (f x) \ b; net \ bot\ \ Im3 l \ b" by (blast intro: tendsto_upperbound [OF tendsto_Im3]) lemma Re_tendsto_lowerbound: "\(f \ l) net; \\<^sub>F x in net. b \ quat.Re (f x); net \ bot\ \ b \ Re l" by (blast intro: tendsto_lowerbound [OF tendsto_Re]) lemma Im1_tendsto_lowerbound: "\(f \ l) net; \\<^sub>F x in net. b \ Im1 (f x); net \ bot\ \ b \ Im1 l" by (blast intro: tendsto_lowerbound [OF tendsto_Im1]) lemma Im2_tendsto_lowerbound: "\(f \ l) net; \\<^sub>F x in net. b \ Im2 (f x); net \ bot\ \ b \ Im2 l" by (blast intro: tendsto_lowerbound [OF tendsto_Im2]) lemma Im3_tendsto_lowerbound: "\(f \ l) net; \\<^sub>F x in net. b \ Im3 (f x); net \ bot\ \ b \ Im3 l" by (blast intro: tendsto_lowerbound [OF tendsto_Im3]) lemma of_real_continuous_iff: "continuous net (\x. quat_of_real(f x)) \ continuous net f" by (simp add: continuous_def tendsto_iff) lemma of_real_continuous_on_iff: "continuous_on S (\x. quat_of_real(f x)) \ continuous_on S f" using continuous_on_Re continuous_on_of_real by fastforce subsection\Quaternions for describing 3D isometries\ subsubsection\The \HIm\ operator\ definition HIm :: "quat \ real^3" where "HIm q \ vector[Im1 q, Im2 q, Im3 q]" lemma HIm_Quat: "HIm (Quat w x y z) = vector[x,y,z]" by (simp add: HIm_def) lemma him_eq: "HIm p = HIm q \ Im1 p = Im1 q \ Im2 p = Im2 q \ Im3 p = Im3 q" by (metis HIm_def vector_3) lemma him_of_real [simp]: "HIm(of_real a) = 0" by (simp add: of_real_eq_Quat HIm_Quat vec_eq_iff vector_def) lemma him_0 [simp]: "HIm 0 = 0" by (metis him_of_real of_real_0) lemma him_1 [simp]: "HIm 1 = 0" by (metis him_of_real of_real_1) lemma him_cnj: "HIm(cnj q) = - HIm q" by (simp add: HIm_def vec_eq_iff vector_def) lemma him_mult_left [simp]: "HIm(of_real a * q) = a *\<^sub>R HIm q" by (simp add: HIm_def vec_eq_iff vector_def) lemma him_mult_right [simp]: "HIm(q * of_real a) = a *\<^sub>R HIm q" by (simp add: HIm_def vec_eq_iff vector_def) lemma him_add [simp]: "HIm(p + q) = HIm p + HIm q" by (simp add: HIm_def vec_eq_iff vector_def) lemma him_minus [simp]: "HIm(-q) = - HIm q" by (simp add: HIm_def vec_eq_iff vector_def) lemma him_diff [simp]: "HIm(p - q) = HIm p - HIm q" by (simp add: HIm_def vec_eq_iff vector_def) lemma him_sum [simp]: "HIm (sum f S) = (\x\S. HIm (f x))" by (induct S rule: infinite_finite_induct) auto lemma linear_him: "linear HIm" by (metis him_add him_mult_right linearI mult.right_neutral mult_scaleR_right of_real_def) subsubsection\The \Hv\ operator\ definition Hv :: "real^3 \ quat" where "Hv v \ Quat 0 (v$1) (v$2) (v$3)" lemma Re_Hv [simp]: "Re(Hv v) = 0" by (simp add: Hv_def) lemma Im1_Hv [simp]: "Im1(Hv v) = v$1" by (simp add: Hv_def) lemma Im2_Hv [simp]: "Im2(Hv v) = v$2" by (simp add: Hv_def) lemma Im3_Hv [simp]: "Im3(Hv v) = v$3" by (simp add: Hv_def) lemma hv_vec: "Hv(vec r) = Quat 0 r r r" by (simp add: Hv_def) lemma hv_eq_zero [simp]: "Hv v = 0 \ v = 0" by (simp add: quat_eq_iff vec_eq_iff) (metis exhaust_3) lemma hv_zero [simp]: "Hv 0 = 0" by simp lemma hv_vector [simp]: "Hv(vector[x,y,z]) = Quat 0 x y z" by (simp add: Hv_def) lemma hv_basis: "Hv(axis 1 1) = \" "Hv(axis 2 1) = \" "Hv(axis 3 1) = \" by (auto simp: Hv_def axis_def quat_ii.code quat_jj.code quat_kk.code) lemma hv_add [simp]: "Hv(x + y) = Hv x + Hv y" by (simp add: Hv_def quat_eq_iff) lemma hv_minus [simp]: "Hv(-x) = -Hv x" by (simp add: Hv_def quat_eq_iff) lemma hv_diff [simp]: "Hv(x - y) = Hv x - Hv y" by (simp add: Hv_def quat_eq_iff) lemma hv_cmult [simp]: "Hv(a *\<^sub>R x) = of_real a * Hv x" by (simp add: Hv_def quat_eq_iff) lemma hv_sum [simp]: "Hv (sum f S) = (\x\S. Hv (f x))" by (induct S rule: infinite_finite_induct) auto lemma hv_inj: "Hv x = Hv y \ x = y" by (simp add: Hv_def quat_eq_iff vec_eq_iff) (metis (full_types) exhaust_3) lemma linear_hv: "linear Hv" by unfold_locales (auto simp: of_real_def) lemma him_hv [simp]: "HIm(Hv x) = x" using HIm_def hv_inj quat_eq_iff by fastforce lemma cnj_hv [simp]: "cnj(Hv v) = -Hv v" using Hv_def cnj.code hv_minus by auto lemma hv_him: "Hv(HIm q) = Quat 0 (Im1 q) (Im2 q) (Im3 q)" by (simp add: HIm_def) lemma hv_him_eq: "Hv(HIm q) = q \ Re q = 0" by (simp add: hv_him quat_eq_iff) lemma dot_hv [simp]: "Hv u \ Hv v = u \ v" by (simp add: Hv_def inner_quat_def inner_vec_def sum_3) lemma norm_hv [simp]: "norm (Hv v) = norm v" by (simp add: norm_eq) subsection\Geometric interpretation of the product of imaginary quaternions\ context includes cross3_syntax begin lemma mult_hv_eq_cross_dot: "Hv x * Hv y = Hv(x \ y) - of_real (x \ y)" by (simp add: quat_eq_iff cross3_simps) text\Representing orthogonal transformations as conjugation or congruence with a quaternion\ lemma orthogonal_transformation_quat_congruence: assumes "norm q = 1" shows "orthogonal_transformation (\x. HIm(cnj q * Hv x * q))" proof - have nq: "(quat.Re q)\<^sup>2 + (Im1 q)\<^sup>2 + (Im2 q)\<^sup>2 + (Im3 q)\<^sup>2 = 1" using assms norm_quat_def by auto have "Vector_Spaces.linear (*\<^sub>R) (*\<^sub>R) (\x. HIm (cnj q * Hv x * q))" proof show "\r b. HIm (cnj q * Hv (r *\<^sub>R b) * q) = r *\<^sub>R HIm (cnj q * Hv b * q)" by (metis him_mult_left hv_cmult mult_scaleR_left mult_scaleR_right scaleR_conv_of_real) qed (simp add: distrib_left distrib_right) moreover have "HIm (cnj q * Hv v * q) \ HIm (cnj q * Hv w * q) = ((quat.Re q)\<^sup>2 + (Im1 q)\<^sup>2 + (Im2 q)\<^sup>2 + (Im3 q)\<^sup>2)\<^sup>2 * (v \ w)" for v w by (simp add: HIm_def inner_vec_def sum_3 power2_eq_square algebra_simps) (*SLOW, like 15s*) ultimately show ?thesis by (simp add: orthogonal_transformation_def linear_def nq) qed lemma orthogonal_transformation_quat_conjugation: assumes "q \ 0" shows "orthogonal_transformation (\x. HIm(inverse q * Hv x * q))" proof - obtain c p where eq: "q = of_real c * p" and 1: "norm p = 1" proof show "q = quat_of_real (norm q) * (inverse (of_real (norm q)) * q)" by (metis assms mult.assoc mult.left_neutral norm_eq_zero of_real_eq_0_iff right_inverse) show "norm (inverse (quat_of_real (norm q)) * q) = 1" using assms by (simp add: norm_mult norm_inverse) qed have "c \ 0" using assms eq by auto then have "HIm (cnj p * Hv x * p) = HIm (inverse (quat_of_real c * p) * Hv x * (quat_of_real c * p))" for x apply (simp add: inverse_mult mult.assoc flip: quat_inverse_eq_cnj [OF 1] of_real_inverse) using quat_of_real_times_commute quat_of_real_times_left_commute quat_of_real_inverse_collapse by (simp add: of_real_def) then show ?thesis using orthogonal_transformation_quat_congruence [OF 1] by (simp add: eq) qed unbundle no_quaternion_syntax end end diff --git a/thys/Tarskis_Geometry/Metric.thy b/thys/Tarskis_Geometry/Metric.thy --- a/thys/Tarskis_Geometry/Metric.thy +++ b/thys/Tarskis_Geometry/Metric.thy @@ -1,64 +1,64 @@ (* Title: Metric and semimetric spaces Author: Tim Makarios , 2012 Maintainer: Tim Makarios *) section "Metric and semimetric spaces" theory Metric -imports "HOL-Analysis.Ordered_Euclidean_Space" +imports "HOL-Analysis.Multivariate_Analysis" begin locale semimetric = fixes dist :: "'p \ 'p \ real" assumes nonneg [simp]: "dist x y \ 0" and eq_0 [simp]: "dist x y = 0 \ x = y" and symm: "dist x y = dist y x" begin lemma refl [simp]: "dist x x = 0" by simp end locale metric = fixes dist :: "'p \ 'p \ real" assumes [simp]: "dist x y = 0 \ x = y" and triangle: "dist x z \ dist y x + dist y z" sublocale metric < semimetric proof { fix w have "dist w w = 0" by simp } note [simp] = this fix x y show "0 \ dist x y" proof - from triangle [of y y x] show "0 \ dist x y" by simp qed show "dist x y = 0 \ x = y" by simp show "dist x y = dist y x" proof - { fix w z have "dist w z \ dist z w" proof - from triangle [of w z z] show "dist w z \ dist z w" by simp qed } hence "dist x y \ dist y x" and "dist y x \ dist x y" by simp+ thus "dist x y = dist y x" by simp qed qed definition norm_dist :: "('a::real_normed_vector) \ 'a \ real" where [simp]: "norm_dist x y \ norm (x - y)" interpretation norm_metric: metric norm_dist proof fix x y show "norm_dist x y = 0 \ x = y" by simp fix z from norm_triangle_ineq [of "x - y" "y - z"] have "norm (x - z) \ norm (x - y) + norm (y - z)" by simp with norm_minus_commute [of x y] show "norm_dist x z \ norm_dist y x + norm_dist y z" by simp qed end diff --git a/thys/Taylor_Models/Float_Topology.thy b/thys/Taylor_Models/Float_Topology.thy --- a/thys/Taylor_Models/Float_Topology.thy +++ b/thys/Taylor_Models/Float_Topology.thy @@ -1,224 +1,224 @@ section \Topology for Floating Point Numbers\ theory Float_Topology imports - "HOL-Analysis.Ordered_Euclidean_Space" + "HOL-Analysis.Multivariate_Analysis" "HOL-Library.Float" begin text \ This topology is totally disconnected and not complete, in which sense is it useful? Perhaps for convergence of intervals?\ unbundle float.lifting instantiation float :: dist begin lift_definition dist_float :: "float \ float \ real" is dist . lemma dist_float_eq_0_iff: "(dist x y = 0) = (x = y)" for x y::float by transfer simp lemma dist_float_triangle2: "dist x y \ dist x z + dist y z" for x y z::float by transfer (rule dist_triangle2) instance .. end instantiation float :: uniformity begin definition uniformity_float :: "(float \ float) filter" where "uniformity_float = (INF e\{0<..}. principal {(x, y). dist x y < e})" instance .. end lemma float_dense_in_real: fixes x :: real assumes "x < y" shows "\r\float. x < r \ r < y" proof - from \x < y\ have "0 < y - x" by simp with reals_Archimedean obtain q' :: nat where q': "inverse (real q') < y - x" and "0 < q'" by blast define q::nat where "q \ 2 ^ nat \bitlen q'\" from bitlen_bounds[of q'] \0 < q'\ have "q' < q" by (auto simp: q_def) then have "inverse q < inverse q'" using \0 < q'\ by (auto simp: divide_simps) with \q' < q\ q' have q: "inverse (real q) < y - x" and "0 < q" by (auto simp: split: if_splits) define p where "p = \y * real q\ - 1" define r where "r = of_int p / real q" from q have "x < y - inverse (real q)" by simp also from \0 < q\ have "y - inverse (real q) \ r" by (simp add: r_def p_def le_divide_eq left_diff_distrib) finally have "x < r" . moreover from \0 < q\ have "r < y" by (simp add: r_def p_def divide_less_eq diff_less_eq less_ceiling_iff [symmetric]) moreover have "r \ float" by (simp add: r_def q_def) ultimately show ?thesis by blast qed lemma real_of_float_dense: fixes x y :: real assumes "x < y" shows "\q :: float. x < real_of_float q \ real_of_float q < y" using float_dense_in_real [OF \x < y\] by (auto elim: real_of_float_cases) instantiation float :: linorder_topology begin definition open_float::"float set \ bool" where "open_float S = (\x\S. \e>0. \y. dist y x < e \ y \ S)" instance proof (standard, intro ext iffI) fix U::"float set" assume "generate_topology (range lessThan \ range greaterThan) U" then show "open U" unfolding open_float_def uniformity_float_def proof (induction U) case UNIV then show ?case by (auto intro!: zero_less_one) next case (Int a b) show ?case proof safe fix x assume "x \ a" "x \ b" with Int(3,4) obtain e1 e2 where "dist (y) (x) < e1 \ y \ a" and "dist (y) (x) < e2 \ y \ b" and "0 < e1" "0 < e2" for y by (auto dest!: bspec) then show "\e>0. \y. dist y x < e \ y \ a \ b" by (auto intro!: exI[where x="min e1 e2"]) qed next case (UN K) show ?case proof safe fix x X assume x: "x \ X" and X: "X \ K" from UN[OF X] x obtain e where "dist (y) (x) < e \ y \ X" "e > 0" for y by auto then show "\e>0. \y. dist (real_of_float y) (real_of_float x) < e \ y \ \K" using x X by (auto intro!: exI[where x=e]) qed next case (Basis s) then show ?case proof safe fix x u::float assume "x < u" then show "\e>0. \y. dist (real_of_float y) (real_of_float x) < e \ y \ {..e>0. \y. dist (real_of_float y) (real_of_float x) < e \ y \ {l<..}" by (force simp add: eventually_principal dist_float_def dist_real_def abs_real_def intro!: exI[where x="(x - l)/2"]) qed qed next fix U::"float set" assume "open U" then obtain e where e: "x \ U \ e x > 0" "x \ U \ dist ( y) ( x) < e x \ y \ U" for x y unfolding open_float_def uniformity_float_def by metis { fix x assume x: "x \ U" obtain e' where e': "e' > 0" "real_of_float e' < e x" using real_of_float_dense[of 0 "e x"] using e(1)[OF x] by auto then have "dist (y) ( x) < e' \ y \ U" for y by (intro e[OF x]) auto then have "\e'>0. \y. dist (y) ( x) < real_of_float e' \ y \ U" using e' by auto } then obtain e' where e': "x \ U \ 0 < e' x" "x \ U \ dist y x < real_of_float (e' x) \ y \ U" for x y by metis then have "U = (\x\U. greaterThan (x - e' x) \ lessThan (x + e' x))" by (auto simp: dist_float_def dist_commute dist_real_def) also have "generate_topology (range lessThan \ range greaterThan) \" by (intro generate_topology_Union generate_topology.Int generate_topology.Basis) auto finally show "generate_topology (range lessThan \ range greaterThan) U" . qed end instance float :: metric_space proof standard fix U::"float set" show "open U = (\x\U. \\<^sub>F (x', y) in uniformity. x' = x \ y \ U)" unfolding open_float_def open_dist uniformity_float_def uniformity_real_def proof safe fix x assume "\x\U. \e>0. \y. dist (real_of_float y) (real_of_float x) < e \ y \ U" "x \ U" then obtain e where "e > 0" "dist (y) (x) < e \ y \ U" for y by auto then show "\\<^sub>F (x', y) in INF e\{0<..}. principal {(x, y). dist x y < e}. x' = x \ y \ U" by (intro eventually_INF1[where i=e]) (auto simp: eventually_principal dist_commute dist_float_def) next fix u assume "\x\U. \\<^sub>F (x', y) in INF e\{0<..}. principal {(x, y). dist x y < e}. x' = x \ y \ U" "u \ U" from this obtain E where E: "E \ {0<..}" "finite E" "\(x', y)\\x\E. {(y', y). dist y' y < x}. x' = u \ y \ U" by (subst (asm) eventually_INF) (auto simp: INF_principal_finite eventually_principal) then show "\e>0. \y. dist (real_of_float y) (real_of_float u) < e \ y \ U" by (intro exI[where x="if E = {} then 1 else Min E"]) (auto simp: dist_commute dist_float_def) qed qed (use dist_float_eq_0_iff dist_float_triangle2 in \auto simp add: uniformity_float_def dist_float_def\) instance float::topological_ab_group_add proof fix a b::float show "((\x. fst x + snd x) \ a + b) (nhds a \\<^sub>F nhds b)" proof (rule tendstoI) fix e::real assume "e > 0" have 1: "(fst \ a) (nhds a \\<^sub>F nhds b)" and 2: "(snd \ b) (nhds a \\<^sub>F nhds b)" by (auto intro!: tendsto_eq_intros filterlim_ident simp: nhds_prod[symmetric]) have "\\<^sub>F x in nhds a \\<^sub>F nhds b. dist (fst x) (a) < e/2" by (rule tendstoD[OF 1]) (use \e > 0\ in auto) moreover have "\\<^sub>F x in nhds a \\<^sub>F nhds b. dist (snd x) (b) < e/2" by (rule tendstoD[OF 2]) (use \e > 0\ in auto) ultimately show "\\<^sub>F x in nhds a \\<^sub>F nhds b. dist (fst x + snd x) (a + b) < e" proof eventually_elim case (elim x) then show ?case by (auto simp: dist_float_def) norm qed qed show "(uminus \ - a) (nhds a)" using filterlim_ident[of "nhds a"] by (auto intro!: tendstoI dest!: tendstoD simp: dist_float_def dist_minus) qed lifting_update float.lifting lifting_forget float.lifting end diff --git a/thys/Transcendence_Series_Hancl_Rucki/Transcendence_Series.thy b/thys/Transcendence_Series_Hancl_Rucki/Transcendence_Series.thy --- a/thys/Transcendence_Series_Hancl_Rucki/Transcendence_Series.thy +++ b/thys/Transcendence_Series_Hancl_Rucki/Transcendence_Series.thy @@ -1,1042 +1,1042 @@ (* Title: Transcendence_Series.thy Authors: Angeliki Koutsoukou-Argyraki and Wenda Li, University of Cambridge *) section \The transcendence of certain infinite series\ theory "Transcendence_Series" imports - "HOL-Analysis.Ordered_Euclidean_Space" + "HOL-Analysis.Multivariate_Analysis" "HOL-Computational_Algebra.Polynomial" Prime_Number_Theorem.Prime_Number_Theorem_Library begin text \ We formalise the proofs of two transcendence criteria by J. Han\v{c}l and P. Rucki that assert the transcendence of the sums of certain infinite series built up by sequences that fulfil certain properties (Theorems 2.1 and 2.2 in \cite{hancl2005}, HanclRucki1 and HanclRucki2 here respectively). Both proofs make use of Roth's celebrated theorem on diophantine approximations to algebraic numbers from 1955 \cite{roth_1955} which we assume and implement within the locale RothsTheorem. A small mistake was detected in the original proof of Theorem 2.1, and the authors suggested to us a fix for the problem (in communication by email). Our formalised proof incorporates this correction (see the Remark in the proof of HanclRucki1). \ subsection \Misc\ lemma powr_less_inverse_iff: fixes x y z::real assumes "x>0""y>0""z>0" shows "x powr y < z \ x < z powr (inverse y)" proof assume "x powr y < z" from powr_less_mono2[OF _ _ this,of "inverse y"] show "x < z powr inverse y" using assms by (auto simp:powr_powr) next assume *:"x < z powr inverse y" from powr_less_mono2[OF _ _ *,of y] show "x powr y < z" using assms by (auto simp:powr_powr) qed lemma powr_less_inverse_iff': fixes x y z::real assumes "x>0""y>0""z>0" shows "z< x powr y \ z powr (inverse y) < x" using powr_less_inverse_iff[symmetric,of _ "inverse y"] assms by auto lemma powr_less_eq_inverse_iff: fixes x y z::real assumes "x>0""y>0""z>0" shows "x powr y \ z \ x \ z powr (inverse y)" by (meson assms(1) assms(2) assms(3) not_less powr_less_inverse_iff') lemma powr_less_eq_inverse_iff': fixes x y z::real assumes "x>0""y>0""z>0" shows "z\ x powr y \ z powr (inverse y) \ x" by (simp add: assms(1) assms(2) assms(3) powr_less_eq_inverse_iff) lemma tendsto_PInfty_mono: assumes "(ereal o f) \ \" "\\<^sub>F x in sequentially. f x \ g x" shows "(ereal o g) \ \" using assms unfolding comp_def tendsto_PInfty_eq_at_top by (elim filterlim_at_top_mono, simp) lemma limsup_infinity_imp_Inf_many: assumes "limsup f = \" shows "(\ m. (\\<^sub>\i. f i > ereal m))" unfolding INFM_nat proof (clarify,rule ccontr) fix m k assume "\ (\n>k. ereal m < f n)" then have "\n>k. f n \ ereal m" by auto then have "\\<^sub>F n in sequentially. f n \ ereal m" using eventually_at_top_dense by blast then have "limsup f \ ereal m" using Limsup_bounded by auto then show False using assms by simp qed lemma snd_quotient_plus_leq: defines "de\(snd o quotient_of)" shows "de (x+y) \ de x * de y " proof - obtain x1 x2 y1 y2 where xy: "quotient_of x = (x1,x2)" "quotient_of y=(y1,y2)" by (meson surj_pair) have "x2>0" "y2>0" using xy quotient_of_denom_pos by blast+ then show ?thesis unfolding de_def comp_def rat_plus_code xy apply (auto split:prod.split simp:Rat.normalize_def Let_def) by (smt div_by_1 gcd_pos_int int_div_less_self mult_eq_0_iff mult_sign_intros(1)) qed lemma quotient_of_inj: "inj quotient_of" unfolding inj_def by (simp add: quotient_of_inject) lemma infinite_inj_imageE: assumes "infinite A" "inj_on f A" "f ` A \ B" shows "infinite B" using assms inj_on_finite by blast lemma incseq_tendsto_limsup: fixes f::"nat \ 'a::{complete_linorder,linorder_topology}" assumes "incseq f" shows "f \ limsup f" using LIMSEQ_SUP assms convergent_def convergent_ereal tendsto_Limsup trivial_limit_sequentially by blast subsection \Main proofs\ text \Since the proof of Roths theorem has not been formalized yet, we implement it into a locale and used it as an assumption.\ locale RothsTheorem = assumes RothsTheorem:"\\ \. algebraic \ \ \ \ \ \ infinite {(p,q). q>0 \ coprime p q \ \\ - of_int p / of_int q\ < 1 / q powr \} \ \ \ 2" theorem (in RothsTheorem) HanclRucki1: fixes a b ::"nat\int" and \ ::real defines "aa\(\n. real_of_int (a n))" and "bb\(\n. real_of_int (b n))" assumes a_pos:"\ k. a k >0" and b_pos:"\ k. b k >0" and "\ >0" and limsup_infy:"limsup (\ k. aa (k+1)/(\i = 0..k. aa i)powr(2+\)*(1/bb (k+1))) = \" and liminf_1:"liminf (\k. aa (k+1) / aa k * bb k / bb (k+1)) > 1" shows "\ algebraic(suminf (\ k. bb k / aa k))" proof - have summable:"summable (\ k. bb k / aa k)" proof (rule ratio_test_convergence) have [simp]:"aa k>0" "bb k>0" for k unfolding aa_def bb_def using a_pos b_pos by auto show "\\<^sub>F n in sequentially. 0 < bb n / aa n" apply (rule eventuallyI) by auto show "1 < liminf (\n. ereal (bb n / aa n / (bb (Suc n) / aa (Suc n))))" using liminf_1 by (auto simp:algebra_simps) qed have [simp]: "aa k>0" "bb k>0" for k unfolding aa_def bb_def by (auto simp add: a_pos b_pos) have ab_1:"aa k\1" "bb k\1" for k unfolding aa_def bb_def using a_pos b_pos by (auto simp add: int_one_le_iff_zero_less) define B where "B=liminf (\x. ereal (aa (x + 1) / aa x * bb x / bb (x + 1)))" define M where "M= (case B of ereal m \ (m+1)/2 | _ \ 2)" have "M > 1" "M < B" using liminf_1 unfolding M_def apply (fold B_def) by (cases B,auto)+ text \ Remark: In the original proof of Theorem 2.1 in \cite{hancl2005} it was claimed in p.534 that from assumption (3) (i.e. @{thm liminf_1}), we obtain that: $\forall A>1~\exists k_0~ \forall k > k_0~ \frac{1}{A} \frac{b_k}{ a_k} > \frac{ b_{k+1}}{ a_{k+1}} $, however note the counterexample where $a_{k+1}= k (a_1 a_2 ... a_k)^{\lceil 2+ \delta \rceil}$ if k is odd, and $a_{k+1} = 2 a_k$ otherwise, with $b_k =1$ for all $k$. In commmunication by email the authors suggested to replace the claim $\forall A>1~\exists k_0~ \forall k > k_0~ \frac{1}{A} \frac{b_k}{ a_k} > \frac{ b_{k+1}}{ a_{k+1}} $ with $\exists A>1~\exists k_0~ \forall k > k_0~ \frac{1}{A} \frac{b_k}{ a_k} > \frac{ b_{k+1}}{ a_{k+1}} $ which solves the problem and the proof proceeds as in the paper. The witness for $\exists A>1 $ is denoted by $M$ here.\ have bb_aa_event:"\\<^sub>F k in sequentially. (1/M)*(bb k / aa k)> bb(k+1)/ aa (k+1)" using less_LiminfD[OF \M < B\[unfolded B_def],simplified] apply eventually_elim using \M > 1\ by (auto simp:divide_simps algebra_simps) have bb_aa_p:"\\<^sub>F k in sequentially. \p. bb(k+p)/ aa (k+p) \ (1/M^p)*(bb k / aa k)" proof - obtain k0 where k0_ineq: "\n\k0. bb (n + 1) / aa (n + 1) < 1 / M * (bb n / aa n)" using bb_aa_event unfolding eventually_sequentially by auto have "bb(k+p)/ aa (k+p) \ (1/M^p)*(bb k / aa k)" when "k\k0" for p k proof (induct p) case 0 then show ?case by auto next case (Suc p) have " bb (k + Suc p) / aa (k + Suc p) < 1 / M * (bb (k+p) / aa (k+p))" using k0_ineq[rule_format,of "k+p"] that by auto also have "... \ 1 / M ^(Suc p) * (bb k / aa k)" using Suc \M>1\ by (auto simp add:field_simps) finally show ?case by auto qed then show ?thesis unfolding eventually_sequentially by auto qed define \ where "\ = suminf (\ k. bb k / aa k)" have \_Inf_many:"\\<^sub>\ k. \\ - (\k = 0..k. bb k / aa k)\ < 1 / prod aa {0..k} powr (2 + \)" proof - have "\\ - (\i = 0..k. bb i / aa i)\ = \\i. bb (i+(k+1)) / aa (i+(k+1))\" for k unfolding \_def apply (subst suminf_minus_initial_segment[of _ "k+1",OF summable]) using atLeast0AtMost lessThan_Suc_atMost by auto moreover have "\\<^sub>\ k. \\i. bb(i+(k+1))/ aa (i+(k+1))\ < 1 / prod aa {0..k} powr (2 + \)" proof - define P where "P = (\ i. \p. bb (i + 1 + p) / aa (i + 1 + p) \ 1 / M ^ p * (bb (i + 1) / aa (i + 1)))" define Q where "Q= (\ i. ereal (M / (M - 1)) < ereal (aa (i + 1) / prod aa {0..i} powr (2 + \) * (1 / bb (i + 1))))" have "\\<^sub>\ i. P i" using bb_aa_p[THEN sequentially_offset, of 1] cofinite_eq_sequentially unfolding P_def by auto moreover have "\\<^sub>\i. Q i" using limsup_infy[THEN limsup_infinity_imp_Inf_many,rule_format,of "(M / (M -1))"] unfolding Q_def . moreover have "\\i. bb(i+(k+1))/ aa (i+(k+1))\ < 1 / prod aa {0..k} powr (2 + \)" when "P k" "Q k" for k proof - have summable_M:"summable (\i. 1 / M ^ i)" apply (rule summable_ratio_test[of "1/M"]) using \M>1\ by auto have "(\i. bb (i + (k + 1)) / aa (i + (k + 1))) \ 0" apply (rule suminf_nonneg) subgoal using summable_ignore_initial_segment[OF summable,of "k+1"] by auto subgoal by (simp add: less_imp_le) done then have "\\i. bb (i + (k + 1)) / aa (i + (k + 1))\ = (\i. bb (i + (k + 1)) / aa (i + (k + 1)))" by auto also have "... \ (\i. 1 / M ^ i * (bb (k + 1) / aa (k + 1)))" apply (rule suminf_le) subgoal using that(1) unfolding P_def by (auto simp add:algebra_simps) subgoal using summable_ignore_initial_segment[OF summable,of "k+1"] by auto subgoal using summable_mult2[OF summable_M,of " bb (k + 1) / aa (k + 1)"] by auto done also have "... = (bb (k + 1) / aa (k + 1)) * (\i. 1 / M ^ i)" using suminf_mult2[OF summable_M,of " bb (k + 1) / aa (k + 1)"] by (auto simp:algebra_simps) also have "... = (bb (k + 1) / aa (k + 1)) * (\i. (1 / M) ^ i)" using \M>1\ by (auto simp: field_simps) also have "... = (bb (k + 1) / aa (k + 1)) * (M / (M -1))" apply (subst suminf_geometric) using \M>1\ by (auto simp: field_simps) also have "... < (bb (k + 1) / aa (k + 1)) * (aa (k + 1) / prod aa {0..k} powr (2 + \) * (1 / bb (k + 1)))" apply (subst mult_less_cancel_left_pos) using that(2) unfolding Q_def by auto also have "... = 1/ prod aa {0..k} powr (2 + \)" using ab_1[of "Suc k"] by auto finally show ?thesis . qed ultimately show ?thesis by (smt INFM_conjI INFM_mono) qed ultimately show ?thesis by auto qed define pq where "pq = (\k. quotient_of (\i=0..k. of_int (b i) / of_int (a i)))" define p q where "p = fst \ pq" and "q = snd \ pq" have coprime_pq:"coprime (p k) (q k)" and q_pos:"q k > 0" and pq_sum:"p k / q k = (\i=0..k. b i / a i)" for k proof - have eq: "quotient_of (\i=0..k. of_int (b i) / of_int (a i)) = (p k, q k)" by (simp add: p_def q_def pq_def) from quotient_of_coprime[OF eq] show "coprime (p k) (q k)" . from quotient_of_denom_pos[OF eq] show "q k > 0" . have "(\i=0..k. b i / a i) = of_rat (\i=0..k. of_int (b i) / of_int (a i))" by (simp add: of_rat_sum of_rat_divide) also have "(\i=0..k. rat_of_int (b i) / rat_of_int (a i)) = rat_of_int (p k) / rat_of_int (q k)" using quotient_of_div[OF eq] by simp finally show "p k / q k = (\i=0..k. b i / a i)" by (simp add:of_rat_divide) qed have \_Inf_many2:"\\<^sub>\ k. \\ - p k / q k\ < 1 / q k powr (2 + \)" using \_Inf_many proof (elim INFM_mono) fix k assume asm:"\\ - (\k = 0..k. bb k / aa k)\ < 1 / prod aa {0..k} powr (2 + \)" have "\\ - real_of_int (p k) / real_of_int (q k)\ = \\ - (\k = 0..k. bb k / aa k)\" using pq_sum unfolding aa_def bb_def by auto also have "... < 1 / prod aa {0..k} powr (2 + \)" using asm by auto also have "... \ 1 / q k powr (2 + \)" proof - have "q k \ prod aa {0..k}" proof (induct k) case 0 then show ?case unfolding q_def pq_def aa_def apply (simp add:rat_divide_code of_int_rat quotient_of_Fract) using ab_1[of 0,unfolded aa_def bb_def] unfolding Let_def normalize_def apply auto by (metis div_by_1 gcd_pos_int less_imp_le less_trans nonneg1_imp_zdiv_pos_iff not_less zdiv_mono2) next case (Suc k) define de where "de=snd \ quotient_of" have "real_of_int (q (Suc k)) = de (\i=0..Suc k. of_int (b i) / of_int (a i))" unfolding q_def pq_def de_def by simp also have "... = de ((\i=0..k. of_int (b i) / of_int (a i)) + of_int (b (Suc k)) / of_int (a (Suc k)))" by simp also have "... \ de (\i=0..k. of_int (b i) / of_int (a i)) * de (of_int (b (Suc k)) / of_int (a (Suc k)))" using snd_quotient_plus_leq[folded de_def] by presburger also have "... = q k * de (of_int (b (Suc k)) / of_int (a (Suc k)))" unfolding q_def pq_def de_def by auto also have "... = q k * snd (Rat.normalize (b (Suc k), a (Suc k)))" by (simp add:rat_divide_code of_int_rat quotient_of_Fract de_def) also have "... \ q k * aa (Suc k)" using ab_1[of "Suc k"] q_pos[of "k"] unfolding normalize_def aa_def bb_def Let_def apply auto by (metis div_by_1 int_one_le_iff_zero_less less_trans nonneg1_imp_zdiv_pos_iff not_less zdiv_mono2 zero_less_one) also have "... \ prod aa {0..k} * aa (Suc k)" using Suc ab_1[of "Suc k"] by auto also have "... = prod aa {0..Suc k}" by (simp add: prod.atLeast0_atMost_Suc) finally show ?case . qed then show ?thesis by (smt \0 < \\ frac_le of_int_0 of_int_le_iff powr_gt_zero powr_mono2 q_pos) qed finally show "\\ - real_of_int (p k) / real_of_int (q k)\ < 1 / real_of_int (q k) powr (2 + \)" . qed define pqs where "pqs={(p, q). q>0 \ coprime p q \ \\ - real_of_int p / real_of_int q\ < 1 / q powr (2 + \)}" have \_infinite:"infinite pqs" proof - define A where "A={k. \\ - (p k) / (q k)\ < 1 / (q k) powr (2 + \)}" have "\\<^sub>\ k. \\ - p k / q k\ < 1 / q k powr (2 + \)" using \_Inf_many2 . then have "infinite A" unfolding Inf_many_def A_def by auto moreover have "inj_on (\k. (p k, q k)) A" proof - define g where "g=(\i. rat_of_int (b i) / rat_of_int (a i))" define f where "f=(\k. \i = 0..k. g i)" have g_pos:"g i>0" for i unfolding g_def by (simp add: a_pos b_pos) have "strict_mono f" unfolding strict_mono_def f_def proof safe fix x y::nat assume "x < y" then have "sum g {0..y} - sum g {0..x} = sum g {x<..y}" apply (subst Groups_Big.sum_diff[symmetric]) by (auto intro:arg_cong2[where f=sum]) also have "... > 0" apply (rule ordered_comm_monoid_add_class.sum_pos) using \x g_pos by auto finally have "sum g {0..y} - sum g {0..x} >0" . then show "sum g {0..x} < sum g {0..y}" by auto qed then have "inj f" using strict_mono_imp_inj_on by auto then have "inj (quotient_of o f)" by (simp add: inj_compose quotient_of_inj) then have "inj (\k. (p k, q k))" unfolding f_def p_def q_def pq_def comp_def apply (fold g_def) by auto then show ?thesis by (auto elim:subset_inj_on) qed moreover have "(\k. (p k, q k)) ` A \ pqs" unfolding A_def pqs_def using coprime_pq q_pos by auto ultimately show ?thesis apply (elim infinite_inj_imageE) by auto qed moreover have "finite pqs" if "\ \ \" proof - obtain m n where \_mn:"\ = (of_int m / of_int n)" and "coprime m n" "n>0" proof - obtain m n where mn:"\\\ = (of_nat m / of_nat n)" "coprime m n" "n\0" using Rats_abs_nat_div_natE[OF \\ \ \\ Rats_abs_nat_div_natE] by metis define m' and n'::int where "m'=(if \ > 0 then nat m else -nat m)" and "n'=nat n" then have "\ = (of_int m' / of_int n')" "coprime m' n'" "n'>0" using mn by auto then show ?thesis using that by auto qed have "pqs \ {(m,n)} \ {x. x \pqs \ - \m\ - 1 \fst x \ fst x \ \m\ + 1 \ 0 snd x < n }" proof (rule subsetI) fix x assume "x \ pqs" define p q where "p=fst x" and "q=snd x" have "q>0" "coprime p q" and pq_less:"\\ - p / q\ < 1 / q powr (2 + \)" using \x\pqs\ unfolding p_def q_def pqs_def by auto have q_lt_n:"qp \ n\q" proof - have "m * q \ n * p" using that \coprime m n\ \coprime p q\ \q>0\ \n>0\ by (metis eq_rat(1) fst_conv int_one_le_iff_zero_less mult.commute normalize_stable not_one_le_zero quotient_of_Fract snd_conv) then have "1/(n*q) \ \m/n - p/q\" using \q>0\ \n>0\ apply (auto simp:field_simps) by (metis add_diff_cancel_left' diff_diff_eq2 diff_zero less_irrefl not_le of_int_diff of_int_lessD of_int_mult) also have "... < 1 / q powr (2 + \)" using pq_less unfolding \_mn by auto also have "... \ 1 / q ^2" proof - have "real_of_int (q\<^sup>2) = q powr 2" apply (subst powr_numeral) unfolding power2_eq_square using \q>0\ by auto also have "... \ q powr (2 + \)" apply (rule powr_mono) using \q>0\ \\>0\ by auto finally have "real_of_int (q\<^sup>2) \ real_of_int q powr (2 + \)" . moreover have "real_of_int q powr (2 + \) > 0" using \0 < q\ by auto ultimately show ?thesis by (auto simp:field_simps) qed finally have " 1 / (n * q) < 1 / q\<^sup>2" . then show ?thesis using \q>0\ \n>0\ unfolding power2_eq_square by (auto simp:field_simps) qed moreover have "- \m\ - 1 \ p \ p \ \m\ + 1" when "m\p \ n\q" proof - define qn where "qn=q/n" have "0m\p \ n\q\] \q>0\ by auto have "\m/n - p / q\ < 1 / q powr (2 + \)" using pq_less unfolding \_mn by simp then have "\p / q - m/n\ < 1 / q powr (2 + \)" by simp then have " m/n- 1 / q powr (2 + \) < p/q \ p/q < m/n + 1 / q powr (2 + \)" unfolding abs_diff_less_iff by auto then have "qn*m- q / q powr (2 + \) < p \ p < qn*m + q / q powr (2 + \)" unfolding qn_def using \q>0\ by (auto simp:field_simps) moreover have "- \m\ - 1 \ qn*m- q / q powr (2 + \)" proof - have "- \m\ \ qn*m" using \0 \qn<1\ apply (cases "m\0") subgoal apply simp by (meson less_eq_real_def mult_nonneg_nonneg neg_le_0_iff_le of_int_0_le_iff order_trans) subgoal by simp done moreover have "- 1 \ - q / q powr (2 + \)" proof - have "q = q powr 1" using \0 < q\ by auto also have "... \q powr (2 + \)" apply (rule powr_mono) using \q>0\ \\>0\ by auto finally have "q \ q powr (2 + \)" . then show ?thesis using \0 < q\ by auto qed ultimately show ?thesis by auto qed moreover have "qn*m + q / q powr (2 + \) \ \m\ + 1" proof - have "qn*m \ \m\" using \0 \qn<1\ apply (cases "m\0") subgoal by (simp add: mult_left_le_one_le) subgoal by (smt of_int_0_le_iff zero_le_mult_iff) done moreover have "q / q powr (2 + \) \ 1" proof - have "q = q powr 1" using \0 < q\ by auto also have "... \q powr (2 + \)" apply (rule powr_mono) using \q>0\ \\>0\ by auto finally have "q \ q powr (2 + \)" . then show ?thesis using \0 < q\ by auto qed ultimately show ?thesis by auto qed ultimately show ?thesis by auto qed ultimately show "x \ {(m, n)} \ {x \ pqs. - \m\ - 1 \ fst x \ fst x \ \m\ + 1 \ 0 < snd x \ snd x < n}" using \x \ pqs\ \q>0\ unfolding p_def q_def by force qed moreover have "finite {x. x \pqs \ - \m\ - 1 \fst x \ fst x \ \m\ + 1 \ 0 snd x < n }" proof - have "finite ({- \m\ - 1..\m\ +1} \ {0<..pqs \ - \m\ - 1 \fst x \ fst x \ \m\ + 1 \ 0 snd x < n } \ ({- \m\ - 1..\m\ +1} \ {0<.._def) using RothsTheorem[rule_format,of \ "2 + \",folded pqs_def] \\ >0\ by auto qed theorem (in RothsTheorem) HanclRucki2: fixes a b ::"nat\int" and \ \ ::real defines "aa\(\n. real_of_int (a n))" and "bb\(\n. real_of_int (b n))" assumes a_pos:"\ k. a k >0" and b_pos:"\ k. b k >0" and "\ >0" and "\ >0" and limsup_infi:"limsup (\ k.(aa (k+1)/(\i = 0..k. aa i)powr(2+(2/\) + \)) * (1/(bb (k+1)))) = \" and ratio_large:"\ k. ( k \ t \ (( aa(k+1)/bb( k+1)) ) powr(1/(1+\)) \ (( aa k/bb k) powr(1/(1+\)))+1)" shows "\ algebraic(suminf (\ k. bb k /aa k)) " proof- have aa_bb_pos[simp]:"aa k>0" "bb k>0" for k unfolding aa_def bb_def using a_pos b_pos by auto have summable:"summable (\ k. bb k / aa k)" proof - define c0 where "c0\(aa t/bb t) powr(1/(1+\)) - t" have ab_k:"(aa k / bb k) powr(1/(1+\)) \ k + c0" when "k\t" for k using that proof (induct k) case 0 then show ?case unfolding c0_def by simp next case (Suc k) have ?case when "\ t\k" proof - have "t = Suc k" using that Suc.prems by linarith then show ?thesis unfolding c0_def by auto qed moreover have ?case when "t \ k" proof - have "(aa(k+1)/bb(k+1)) powr(1/(1+\)) \ ( aa k/bb k) powr(1/(1+\))+1" using ratio_large[rule_format,OF that] by blast then show ?thesis using Suc(1)[OF that] by simp qed ultimately show ?case by auto qed have "summable (\k. 1 / (k + c0) powr (1+\))" proof - have "c0 + t > 0" unfolding c0_def using aa_bb_pos[of t] by (simp,linarith) then have "summable (\k. 1 / (k + (c0+t)) powr (1+\))" using summable_hurwitz_zeta_real[of "1+\" "c0+t"] apply (subst (asm) powr_minus_divide) using \\>0\ by auto then show ?thesis apply (rule_tac summable_offset[of _ t]) by (auto simp:algebra_simps) qed moreover have "bb k / aa k \ 1 / (k + c0) powr (1+\)" when "k\t" for k proof - have "(aa t / bb t) powr (1 / (1 + \)) > 0" apply simp by (metis \\k. 0 < aa k\ \\k. 0 < bb k\ less_numeral_extra(3)) then have "k + c0 >0" unfolding c0_def using that by linarith then have "aa k / bb k \ (k + c0) powr (1+\)" using ab_k[OF that] apply (subst (asm) powr_less_eq_inverse_iff') using \\ >0\ by auto then have "inverse (aa k / bb k) \ inverse ((k + c0) powr (1+\))" apply (elim le_imp_inverse_le) using \k + c0 >0\ by auto then show ?thesis by (simp add: inverse_eq_divide) qed ultimately show ?thesis apply (elim summable_comparison_test'[where N=t]) using aa_bb_pos by (simp add: less_eq_real_def) qed have 7:"\\<^sub>\ k. 1 / (M powr (\/(1+\)) * (\i = 0..k. aa i) powr(2+ \ * \ / (1+\))) > (bb (k+1) / aa(k+1)) powr (\ / (1+\))" when "M > 0" for M proof - define tt where "tt= (\i. prod aa {0..i} powr (2 + 2 / \ + \))" have tt_pos:"tt i > 0" for i unfolding tt_def apply(subst powr_gt_zero,induct i) subgoal by (metis aa_bb_pos(1) order_less_irrefl prod_pos) subgoal by (metis \\k. 0 < aa k\ order_less_irrefl prod_pos) done have "\\<^sub>\i. M < (aa (i + 1) / tt i * (1 / bb (i + 1)))" using limsup_infinity_imp_Inf_many[OF limsup_infi,rule_format,of M] unfolding tt_def by auto then have "\\<^sub>\i. 1 / (M * tt i) > (bb (i+1) / aa (i+1))" apply (elim INFM_mono) using \M>0\ tt_pos by (auto simp:divide_simps algebra_simps) then have "\\<^sub>\i. (1 / (M * tt i)) powr (\/(1+\)) > (bb (i+1) / aa (i+1)) powr (\/(1+\))" apply (elim INFM_mono powr_less_mono2[rotated 2]) by (simp_all add: assms(6) pos_add_strict less_eq_real_def) moreover have "tt i powr (\/(1+\)) = prod aa {0..i} powr (2 + \ * \ / (1+\))" for i unfolding tt_def apply (auto simp:powr_powr) using \\>0\ by (simp add:divide_simps,simp add:algebra_simps) ultimately show ?thesis apply (elim INFM_mono) by (smt nonzero_mult_div_cancel_left powr_divide powr_mult powr_one_eq_one that tt_pos zero_less_divide_iff) qed have 8:"\\<^sub>F k in sequentially. \s. (( aa(k+s)/bb( k+s))) \ ((( aa k/bb k) powr(1/(1+\))) +s)powr(1+\)" using eventually_ge_at_top[of t] proof eventually_elim case (elim k) define ff where "ff=(\k. (aa k / bb k) powr (1 / (1 + \)))" have 11:"ff k+s \ ff (k+s)" for s proof (induct s) case 0 then show ?case unfolding ff_def by auto next case (Suc s) then have "ff k + Suc s \ ff (k+Suc s)" using ratio_large[rule_format,of "k+s"] \ t \ k\ unfolding ff_def by auto then show ?case by simp qed then have "(ff k+s) powr (1+\) \ ff (k+s) powr (1+\)" for s apply (rule powr_mono2[of "1+\",rotated 2]) unfolding ff_def using \\>0\ by auto then show ?case unfolding ff_def using \\>0\ apply (auto simp add:powr_powr) by (simp add: a_pos aa_def b_pos bb_def le_less) qed have 9: "(\r. 1/((z+real r)powr(1+\))) \ 1/(\ *(z-1) powr \)" "summable (\i. 1/((z+ real i)powr(1+\)))" when "z>1" for z proof - define f where "f= (\r. 1/((z+ r)powr(1+\)))" have "((\x. f (x - 1)) has_integral ((z-1) powr - \) / \) {0..}" proof - have "((\x. (z-1 + x) powr (- 1 - \)) has_integral ((z-1) powr - \) / \) {0<..}" using powr_has_integral_at_top[of 0 "z-1" "- 1 - \",simplified] \z>1\ \\>0\ by auto then have "((\x. (z-1 + x) powr (- 1 - \)) has_integral ((z-1) powr - \) / \) {0..}" apply (subst (asm) has_integral_closure[symmetric]) by (auto simp add: negligible_convex_frontier) then show ?thesis apply (rule has_integral_cong[THEN iffD1, rotated 1]) unfolding f_def by (smt powr_minus_divide) qed moreover have "\x. 0 \ x \ 0 \ f (x - 1)" unfolding f_def by simp moreover have "\x y. 0 \ x \ x \ y \ f (y - 1) \ f (x - 1)" unfolding f_def by (smt assms(6) frac_le powr_mono2 powr_nonneg_iff that) ultimately have "summable (\i. f (real i))" "(\i. f (real i)) \ (z - 1) powr - \ / \" using decreasing_sum_le_integral[of "\x. f (x-1)" "((z-1) powr - \) / \",simplified] by auto moreover have "(z - 1) powr - \ / \ = 1/(\ *(z-1) powr \)" by (simp add: powr_minus_divide) ultimately show "(\i. f (real i)) \ 1/(\ *(z-1) powr \)" by auto show "summable (\i. f (real i))" using \summable (\i. f (real i))\ . qed have u:"(\k.( aa k / bb k)) \ \" proof - define ff where "ff\(\x. ereal (aa x / bb x))" have "limsup ff = \" proof - define tt where "tt= (\i. prod aa {0..i} powr (2 + 2 / \ + \))" have "tt i \ 1" for i unfolding tt_def apply (rule ge_one_powr_ge_zero) subgoal apply (rule linordered_nonzero_semiring_class.prod_ge_1) by (simp add: a_pos aa_def int_one_le_iff_zero_less) subgoal by (simp add: \\>0\ \\>0\ less_imp_le) done then have "limsup (\x. (aa (x + 1) / tt x * (1 / bb (x + 1)))) \ limsup (\x. aa (x + 1) / bb (x + 1))" apply (intro Limsup_mono eventuallyI) apply (auto simp add:field_simps order.order_iff_strict) by (metis aa_bb_pos(1) div_by_1 frac_less2 less_irrefl less_numeral_extra(1) not_le) also have "... = limsup (\x. aa x / bb x)" by (subst limsup_shift,simp) finally have "limsup (\x. ereal (aa (x + 1) / tt x * (1 / bb (x + 1)))) \ limsup (\x. ereal (aa x / bb x))" . moreover have "limsup (\x. ereal (aa (x + 1) / tt x * (1 / bb (x + 1)))) = \" using limsup_infi unfolding tt_def by auto ultimately show ?thesis unfolding ff_def using ereal_infty_less_eq2(1) by blast qed then have "limsup (\k. ff (k+t)) = \" by (simp add: limsup_shift_k) moreover have "incseq (\k. ff (k+t))" proof (rule incseq_SucI) fix k::nat define gg where "gg\(\x. (aa x / bb x))" have "(gg (k+t)) powr (1 / (1 + \)) + 1 \ (gg (Suc (k+t))) powr (1 / (1 + \))" using ratio_large[rule_format, of "k+t",simplified] unfolding gg_def by auto then have "(gg (k+t)) powr (1 / (1 + \)) \ (gg (Suc (k+t))) powr (1 / (1 + \))" by auto then have "gg (k+t)\ gg (Suc (k+t))" by (smt aa_bb_pos(1) aa_bb_pos(2) assms(6) divide_pos_pos gg_def powr_less_mono2) then show "ff (k + t) \ ff (Suc k + t)" unfolding gg_def ff_def by auto qed ultimately have "(\k. ff (k+t)) \ \" using incseq_tendsto_limsup by fastforce then show ?thesis unfolding ff_def unfolding tendsto_def apply (subst eventually_sequentially_seg[symmetric,of _ t]) by simp qed define \ where "\ = suminf (\ k. bb k / aa k)" have 10:"\\<^sub>F k in sequentially. \\ - (\k = 0..k. bb k / aa k)\ < 2 / \ * (bb (k+1) / aa (k+1)) powr (\ / (1+\))" using 8[THEN sequentially_offset,of 1] eventually_ge_at_top[of t] u[unfolded tendsto_PInfty,rule_format,THEN sequentially_offset ,of "(2 powr (1/\) / (2 powr (1/\) - 1)) powr (1+\)" 1] proof eventually_elim case (elim k) define tt where "tt=(aa (k + 1) / bb (k + 1)) powr (1 / (1 + \))" have "tt>1" proof - have "(aa k / bb k) powr (1 / (1 + \)) > 0" by (metis a_pos aa_def b_pos bb_def divide_eq_0_iff less_irrefl of_int_eq_0_iff powr_gt_zero) then show ?thesis using ratio_large[rule_format,OF \k\t\] unfolding tt_def by argo qed have "\\ - (\i = 0..k. bb i / aa i)\ = \\i. bb (i+(k+1)) / aa (i+(k+1))\" unfolding \_def apply (subst suminf_minus_initial_segment[of _ "k+1",OF summable]) using atLeast0AtMost lessThan_Suc_atMost by auto also have "... = (\i. bb (i+(k+1)) / aa (i+(k+1)))" proof - have "(\i. bb (i+(k+1)) / aa (i+(k+1))) > 0" apply (rule suminf_pos) subgoal using summable[THEN summable_ignore_initial_segment,of "k+1"] . subgoal by (simp add: a_pos aa_def b_pos bb_def) done then show ?thesis by auto qed also have "... \ (\i. 1 / (tt + i) powr (1 + \))" proof (rule suminf_le) define ff where "ff=(\k n. (aa (k+1) / bb (k+1)) powr (1 / (1 + \)) + real n)" have "bb (n + (k + 1)) / aa (n + (k + 1)) \ 1 / (ff k n) powr (1 + \)" for n proof - have "ff k n powr (1 + \) \ aa (k + n +1 ) / bb (k + n+1 )" using elim(1)[rule_format,of n] unfolding ff_def by auto moreover have "ff k n powr (1 + \) > 0" unfolding ff_def by (smt elim(2) of_nat_0_le_iff powr_gt_zero ratio_large) ultimately have "1 / ff k n powr (1 + \) \ bb (k + (n + 1)) / aa (k + (n + 1))" apply (drule_tac le_imp_inverse_le) by (auto simp add: inverse_eq_divide) then show ?thesis by (auto simp:algebra_simps) qed then show " \n. bb (n + (k + 1)) / aa (n + (k + 1)) \ 1 / (tt + real n) powr (1 + \)" unfolding ff_def tt_def by auto show "summable (\i. bb (i + (k + 1)) / aa (i + (k + 1)))" using summable[THEN summable_ignore_initial_segment,of "k+1"] . show "summable (\x. 1 / (tt + real x) powr (1 + \))" using 9(2)[OF \tt>1\] . qed also have "... \ 1/(\ *(tt-1) powr \)" using 9[OF \tt>1\] by simp also have "... < 2 / (\ *tt powr \)" proof - have "((2 powr (1/\) / (2 powr (1/\) - 1)) powr (1 + \)) < (aa (k+1) / bb (k+1))" using elim(3) by auto then have "2 powr (1/\) / (2 powr (1/\) - 1) < tt" unfolding tt_def apply (drule_tac powr_less_mono2[rotated 2,where a="1/ (1 + \)"]) using \\>0\ apply (auto simp:powr_powr ) by (subst (asm) powr_one,auto simp add:field_simps) then have " tt < (tt-1) * (2 powr (1/\))" using \\>0\ by (auto simp:divide_simps algebra_simps) then have "tt powr \ < 2 * (tt - 1) powr \" apply (drule_tac powr_less_mono2[rotated 2,where a="\"]) using \\>0\ \tt>1\ by (auto simp:powr_powr powr_mult) then show ?thesis using \\>0\ \tt>1\ by (auto simp:divide_simps) qed also have "... = 2 / \ * (bb (k + 1) / aa (k + 1)) powr (\ / (1 + \))" unfolding tt_def using \\>0\ by (auto simp:powr_powr divide_simps algebra_simps powr_divide) finally show ?case . qed define pq where "pq = (\k. quotient_of (\i=0..k. of_int (b i) / of_int (a i)))" define p q where "p = fst \ pq" and "q = snd \ pq" have coprime_pq:"coprime (p k) (q k)" and q_pos:"q k > 0" and pq_sum:"p k / q k = (\i=0..k. b i / a i)" for k proof - have eq: "quotient_of (\i=0..k. of_int (b i) / of_int (a i)) = (p k, q k)" by (simp add: p_def q_def pq_def) from quotient_of_coprime[OF eq] show "coprime (p k) (q k)" . from quotient_of_denom_pos[OF eq] show "q k > 0" . have "(\i=0..k. b i / a i) = of_rat (\i=0..k. of_int (b i) / of_int (a i))" by (simp add: of_rat_sum of_rat_divide) also have "(\i=0..k. rat_of_int (b i) / rat_of_int (a i)) = rat_of_int (p k) / rat_of_int (q k)" using quotient_of_div[OF eq] by simp finally show "p k / q k = (\i=0..k. b i / a i)" by (simp add:of_rat_divide) qed have \_Inf_many:"\\<^sub>\ k. \\ - p k / q k\ < 1 / q k powr (2 + \ * \ / (1 + \))" proof - have *:"\\<^sub>\k. (bb (Suc k) / aa (Suc k)) powr (\ / (1 + \)) < \ / (2 * prod aa {0..k} powr (2 + \ * \ / (1 + \)))" using 7[of "(2 / \) powr ((1+\)/ \)"] \\>0\ by (auto simp:powr_powr) have **:"\\<^sub>\k. \\ - (\k = 0..k. bb k / aa k)\ < 2 / \ * (bb (k + 1) / aa (k + 1)) powr (\ / (1 + \))" using 10[unfolded cofinite_eq_sequentially[symmetric]] . from INFM_conjI[OF * **] show ?thesis proof (elim INFM_mono) fix k assume asm:"(bb (Suc k) / aa (Suc k)) powr (\ / (1 + \)) < \ / (2 * prod aa {0..k} powr (2 + \ * \ / (1 + \))) \ \\ - (\k = 0..k. bb k / aa k)\ < 2 / \ * (bb (k + 1) / aa (k + 1)) powr (\ / (1 + \))" have "\\ - real_of_int (p k) / real_of_int (q k)\ = \\ - (\k = 0..k. bb k / aa k)\" using pq_sum unfolding aa_def bb_def by auto also have "... < (2 / \) * (bb (k+1) / aa (k+1)) powr (\ / (1+\))" using asm by auto also have "... < (2 / \) * (\ / (2 * prod aa {0..k} powr (2 + \ * \ / (1 + \))))" apply (rule mult_strict_left_mono) using asm \\>0\ by auto also have "... = 1/ prod aa {0..k} powr (2 + \ * \ / (1 + \))" using \\>0\ by auto also have "... \ 1 / q k powr (2 + \ * \ / (1 + \))" proof - have "q k \ prod aa {0..k}" proof (induct k) case 0 then show ?case unfolding q_def pq_def aa_def apply (simp add:rat_divide_code of_int_rat quotient_of_Fract) using aa_bb_pos[of 0,unfolded aa_def bb_def] unfolding Let_def normalize_def apply auto by (metis div_by_1 less_imp_le less_trans nonneg1_imp_zdiv_pos_iff not_less zdiv_mono2) next case (Suc k) define de where "de=snd \ quotient_of" have "real_of_int (q (Suc k)) = de (\i=0..Suc k. of_int (b i) / of_int (a i))" unfolding q_def pq_def de_def by simp also have "... = de ((\i=0..k. of_int (b i) / of_int (a i)) + of_int (b (Suc k)) / of_int (a (Suc k)))" by simp also have "... \ de (\i=0..k. of_int (b i) / of_int (a i)) * de (of_int (b (Suc k)) / of_int (a (Suc k)))" using snd_quotient_plus_leq[folded de_def] by presburger also have "... = q k * de (of_int (b (Suc k)) / of_int (a (Suc k)))" unfolding q_def pq_def de_def by auto also have "... = q k * snd (Rat.normalize (b (Suc k), a (Suc k)))" by (simp add:rat_divide_code of_int_rat quotient_of_Fract de_def) also have "... \ q k * aa (Suc k)" using aa_bb_pos[of "Suc k"] q_pos[of "k"] unfolding normalize_def aa_def bb_def Let_def apply auto by (metis div_by_1 int_one_le_iff_zero_less less_trans nonneg1_imp_zdiv_pos_iff not_less zdiv_mono2 zero_less_one) also have "... \ prod aa {0..k} * aa (Suc k)" using Suc aa_bb_pos[of "Suc k"] by auto also have "... = prod aa {0..Suc k}" by (simp add: prod.atLeast0_atMost_Suc) finally show ?case . qed then show ?thesis apply (rule_tac divide_left_mono) apply (rule powr_mono2) using \\>0\ \\>0\ q_pos[of k] apply (auto simp:powr_mult[symmetric]) by (metis aa_bb_pos(1) less_irrefl) qed finally show "\\ - p k / q k\ < 1 / q k powr (2 + \ * \ / (1 + \))" . qed qed define pqs where "pqs={(p, q). q>0 \ coprime p q \ \\ - real_of_int p / real_of_int q\ < 1 / q powr (2 + \ * \ / (1 + \))}" have \_infinite:"infinite pqs" proof - define A where "A={k. \\ - (p k) / (q k)\ < 1 / (q k) powr (2 + \ * \ / (1 + \))}" note \_Inf_many then have "infinite A" unfolding Inf_many_def A_def by auto moreover have "inj_on (\k. (p k, q k)) A" proof - define g where "g=(\i. rat_of_int (b i) / rat_of_int (a i))" define f where "f=(\k. \i = 0..k. g i)" have g_pos:"g i>0" for i unfolding g_def by (simp add: a_pos b_pos) have "strict_mono f" unfolding strict_mono_def f_def proof safe fix x y::nat assume "x < y" then have "sum g {0..y} - sum g {0..x} = sum g {x<..y}" apply (subst Groups_Big.sum_diff[symmetric]) by (auto intro:arg_cong2[where f=sum]) also have "... > 0" apply (rule ordered_comm_monoid_add_class.sum_pos) using \x g_pos by auto finally have "sum g {0..y} - sum g {0..x} >0" . then show "sum g {0..x} < sum g {0..y}" by auto qed then have "inj f" using strict_mono_imp_inj_on by auto then have "inj (quotient_of o f)" by (simp add: inj_compose quotient_of_inj) then have "inj (\k. (p k, q k))" unfolding f_def p_def q_def pq_def comp_def apply (fold g_def) by auto then show ?thesis by (auto elim:subset_inj_on) qed moreover have "(\k. (p k, q k)) ` A \ pqs" unfolding A_def pqs_def using coprime_pq q_pos by auto ultimately show ?thesis apply (elim infinite_inj_imageE) by auto qed moreover have "finite pqs" if "\ \ \" proof - obtain m n where \_mn:"\ = (of_int m / of_int n)" and "coprime m n" "n>0" proof - obtain m n where mn:"\\\ = (of_nat m / of_nat n)" "coprime m n" "n\0" using Rats_abs_nat_div_natE[OF \\ \ \\ Rats_abs_nat_div_natE] by metis define m' and n'::int where "m'=(if \ > 0 then nat m else -nat m)" and "n'=nat n" then have "\ = (of_int m' / of_int n')" "coprime m' n'" "n'>0" using mn by auto then show ?thesis using that by auto qed have "pqs \ {(m,n)} \ {x. x \pqs \ - \m\ - 1 \fst x \ fst x \ \m\ + 1 \ 0 snd x < n }" proof (rule subsetI) fix x assume "x \ pqs" define p q where "p=fst x" and "q=snd x" have "q>0" "coprime p q" and pq_less:"\\ - p / q\ < 1 / q powr (2 + \ * \ / (1 + \))" using \x\pqs\ unfolding p_def q_def pqs_def by auto have q_lt_n:"qp \ n\q" proof - have "m * q \ n * p" using that \coprime m n\ \coprime p q\ \q>0\ \n>0\ by (metis eq_rat(1) fst_conv int_one_le_iff_zero_less mult.commute normalize_stable not_one_le_zero quotient_of_Fract snd_conv) then have "1/(n*q) \ \m/n - p/q\" using \q>0\ \n>0\ apply (auto simp:field_simps) by (metis add_diff_cancel_left' diff_diff_eq2 diff_zero less_irrefl not_le of_int_diff of_int_lessD of_int_mult) also have "... < 1 / q powr (2 + \ * \ / (1 + \))" using pq_less unfolding \_mn by auto also have "... \ 1 / q ^2" proof - have "real_of_int (q\<^sup>2) = q powr 2" apply (subst powr_numeral) unfolding power2_eq_square using \q>0\ by auto also have "... \ q powr (2 + \ * \ / (1 + \))" apply (rule powr_mono) using \q>0\ \\>0\ \\>0\ by auto finally have "real_of_int (q\<^sup>2) \ real_of_int q powr (2 + \ * \ / (1 + \))" . moreover have "real_of_int q powr (2 + \ * \ / (1 + \)) > 0" using \0 < q\ by auto ultimately show ?thesis by (auto simp:field_simps) qed finally have " 1 / (n * q) < 1 / q\<^sup>2" . then show ?thesis using \q>0\ \n>0\ unfolding power2_eq_square by (auto simp:field_simps) qed moreover have "- \m\ - 1 \ p \ p \ \m\ + 1" when "m\p \ n\q" proof - define qn where "qn=q/n" have "0m\p \ n\q\] \q>0\ by auto have "\m/n - p / q\ < 1 / q powr (2 + \ * \ / (1 + \))" using pq_less unfolding \_mn by simp then have "\p / q - m/n\ < 1 / q powr (2 + \ * \ / (1 + \))" by simp then have " m/n- 1 / q powr (2 + \ * \ / (1 + \)) < p/q \ p/q < m/n + 1 / q powr (2 + \ * \ / (1 + \))" unfolding abs_diff_less_iff by auto then have "qn*m- q / q powr (2 + \ * \ / (1 + \)) < p \ p < qn*m + q / q powr (2 + \ * \ / (1 + \))" unfolding qn_def using \q>0\ by (auto simp:field_simps) moreover have "- \m\ - 1 \ qn*m- q / q powr (2 + \ * \ / (1 + \))" proof - have "- \m\ \ qn*m" using \0 \qn<1\ apply (cases "m\0") subgoal apply simp by (meson less_eq_real_def mult_nonneg_nonneg neg_le_0_iff_le of_int_0_le_iff order_trans) subgoal by simp done moreover have "- 1 \ - q / q powr (2 + \ * \ / (1 + \))" proof - have "q = q powr 1" using \0 < q\ by auto also have "... \q powr (2 + \ * \ / (1 + \))" apply (rule powr_mono) using \q>0\ \\>0\ \\>0\ by auto finally have "q \ q powr (2 + \ * \ / (1 + \))" . then show ?thesis using \0 < q\ by auto qed ultimately show ?thesis by auto qed moreover have "qn*m + q / q powr (2 + \ * \ / (1 + \)) \ \m\ + 1" proof - have "qn*m \ \m\" using \0 \qn<1\ apply (cases "m\0") subgoal by (simp add: mult_left_le_one_le) subgoal by (smt of_int_0_le_iff zero_le_mult_iff) done moreover have "q / q powr (2 + \ * \ / (1 + \)) \ 1" proof - have "q = q powr 1" using \0 < q\ by auto also have "... \q powr (2 + \ * \ / (1 + \))" apply (rule powr_mono) using \q>0\ \\>0\ \\>0\ by auto finally have "q \ q powr (2 + \ * \ / (1 + \))" . then show ?thesis using \0 < q\ by auto qed ultimately show ?thesis by auto qed ultimately show ?thesis by auto qed ultimately show "x \ {(m, n)} \ {x \ pqs. - \m\ - 1 \ fst x \ fst x \ \m\ + 1 \ 0 < snd x \ snd x < n}" using \x \ pqs\ \q>0\ unfolding p_def q_def by force qed moreover have "finite {x. x \pqs \ - \m\ - 1 \fst x \ fst x \ \m\ + 1 \ 0 snd x < n }" proof - have "finite ({- \m\ - 1..\m\ +1} \ {0<..pqs \ - \m\ - 1 \fst x \ fst x \ \m\ + 1 \ 0 snd x < n } \ ({- \m\ - 1..\m\ +1} \ {0<.._def) using RothsTheorem[rule_format,of \ "2 + \ * \ / (1 + \)",folded pqs_def] \\ >0\ \\>0\ apply (auto simp:divide_simps ) by (meson mult_le_0_iff not_less) qed subsection\ Acknowledgements\ text\A.K.-A. and W.L. were supported by the ERC Advanced Grant ALEXANDRIA (Project 742178) funded by the European Research Council and led by Professor Lawrence Paulson at the University of Cambridge, UK. Thanks to Iosif Pinelis for his clarification on MathOverflow regarding the summmability of the series in @{thm [source] RothsTheorem.HanclRucki2} @{url "https://mathoverflow.net/questions/323069/why-is-this-series-summable"} and to Manuel Eberl for his helpful comments.\ end \ No newline at end of file diff --git a/thys/Triangle/Angles.thy b/thys/Triangle/Angles.thy --- a/thys/Triangle/Angles.thy +++ b/thys/Triangle/Angles.thy @@ -1,253 +1,253 @@ (* File: Angles.thy Author: Manuel Eberl Definition of angles between vectors and between three points. *) section \Definition of angles\ theory Angles imports - "HOL-Analysis.Ordered_Euclidean_Space" + "HOL-Analysis.Multivariate_Analysis" begin lemma collinear_translate_iff: "collinear (((+) a) ` A) \ collinear A" by (auto simp: collinear_def) definition vangle where "vangle u v = (if u = 0 \ v = 0 then pi / 2 else arccos (u \ v / (norm u * norm v)))" definition angle where "angle a b c = vangle (a - b) (c - b)" lemma angle_altdef: "angle a b c = arccos ((a - b) \ (c - b) / (dist a b * dist c b))" by (simp add: angle_def vangle_def dist_norm) lemma vangle_0_left [simp]: "vangle 0 v = pi / 2" and vangle_0_right [simp]: "vangle u 0 = pi / 2" by (simp_all add: vangle_def) lemma vangle_refl [simp]: "u \ 0 \ vangle u u = 0" by (simp add: vangle_def dot_square_norm power2_eq_square) lemma angle_refl [simp]: "angle a a b = pi / 2" "angle a b b = pi / 2" by (simp_all add: angle_def) lemma angle_refl_mid [simp]: "a \ b \ angle a b a = 0" by (simp add: angle_def) lemma cos_vangle: "cos (vangle u v) = u \ v / (norm u * norm v)" unfolding vangle_def using Cauchy_Schwarz_ineq2[of u v] by (auto simp: field_simps) lemma cos_angle: "cos (angle a b c) = (a - b) \ (c - b) / (dist a b * dist c b)" by (simp add: angle_def cos_vangle dist_norm) lemma inner_conv_angle: "(a - b) \ (c - b) = dist a b * dist c b * cos (angle a b c)" by (simp add: cos_angle) lemma vangle_commute: "vangle u v = vangle v u" by (simp add: vangle_def inner_commute mult.commute) lemma angle_commute: "angle a b c = angle c b a" by (simp add: angle_def vangle_commute) lemma vangle_nonneg: "vangle u v \ 0" and vangle_le_pi: "vangle u v \ pi" using Cauchy_Schwarz_ineq2[of u v] by (auto simp: vangle_def field_simps intro!: arccos_lbound arccos_ubound) lemmas vangle_bounds = vangle_nonneg vangle_le_pi lemma angle_nonneg: "angle a b c \ 0" and angle_le_pi: "angle a b c \ pi" using vangle_bounds unfolding angle_def by blast+ lemmas angle_bounds = angle_nonneg angle_le_pi lemma sin_vangle_nonneg: "sin (vangle u v) \ 0" using vangle_bounds by (rule sin_ge_zero) lemma sin_angle_nonneg: "sin (angle a b c) \ 0" using angle_bounds by (rule sin_ge_zero) lemma vangle_eq_0D: assumes "vangle u v = 0" shows "norm u *\<^sub>R v = norm v *\<^sub>R u" proof - from assms have "u \ v = norm u * norm v" using arccos_eq_iff[of "(u \ v) / (norm u * norm v)" 1] Cauchy_Schwarz_ineq2[of u v] by (fastforce simp: vangle_def split: if_split_asm) thus ?thesis by (subst (asm) norm_cauchy_schwarz_eq) simp_all qed lemma vangle_eq_piD: assumes "vangle u v = pi" shows "norm u *\<^sub>R v + norm v *\<^sub>R u = 0" proof - from assms have "(-u) \ v = norm (-u) * norm v" using arccos_eq_iff[of "(u \ v) / (norm u * norm v)" "-1"] Cauchy_Schwarz_ineq2[of u v] by (simp add: field_simps vangle_def split: if_split_asm) thus ?thesis by (subst (asm) norm_cauchy_schwarz_eq) simp_all qed lemma dist_triangle_eq: fixes a b c :: "'a :: real_inner" shows "(dist a c = dist a b + dist b c) \ dist a b *\<^sub>R (c - b) + dist b c *\<^sub>R (a - b) = 0" using norm_triangle_eq[of "b - a" "c - b"] by (simp add: dist_norm norm_minus_commute algebra_simps) lemma angle_eq_pi_imp_dist_additive: assumes "angle a b c = pi" shows "dist a c = dist a b + dist b c" using vangle_eq_piD[OF assms[unfolded angle_def]] by (subst dist_triangle_eq) (simp add: dist_norm norm_minus_commute) lemma orthogonal_iff_vangle: "orthogonal u v \ vangle u v = pi / 2" using arccos_eq_iff[of "u \ v / (norm u * norm v)" 0] Cauchy_Schwarz_ineq2[of u v] by (auto simp: vangle_def orthogonal_def) lemma cos_minus1_imp_pi: assumes "cos x = -1" "x \ 0" "x < 3 * pi" shows "x = pi" proof - have "cos (x - pi) = 1" by (simp add: assms) then obtain n :: int where n: "of_int n = (x / pi - 1) / 2" by (subst (asm) cos_one_2pi_int) (auto simp: field_simps) also from assms have "\ \ {-1<..<1}" by (auto simp: field_simps) finally have "n = 0" by simp with n show ?thesis by simp qed lemma vangle_eqI: assumes "u \ 0" "v \ 0" "w \ 0" "x \ 0" assumes "(u \ v) * norm w * norm x = (w \ x) * norm u * norm v" shows "vangle u v = vangle w x" using assms Cauchy_Schwarz_ineq2[of u v] Cauchy_Schwarz_ineq2[of w x] unfolding vangle_def by (auto simp: arccos_eq_iff field_simps) lemma angle_eqI: assumes "a \ b" "a \ c" "d \ e" "d \ f" assumes "((b-a) \ (c-a)) * dist d e * dist d f = ((e-d) \ (f-d)) * dist a b * dist a c" shows "angle b a c = angle e d f" using assms unfolding angle_def by (intro vangle_eqI) (simp_all add: dist_norm norm_minus_commute) lemma cos_vangle_eqD: "cos (vangle u v) = cos (vangle w x) \ vangle u v = vangle w x" by (rule cos_inj_pi) (simp_all add: vangle_bounds) lemma cos_angle_eqD: "cos (angle a b c) = cos (angle d e f) \ angle a b c = angle d e f" unfolding angle_def by (rule cos_vangle_eqD) lemma sin_vangle_zero_iff: "sin (vangle u v) = 0 \ vangle u v \ {0, pi}" proof assume "sin (vangle u v) = 0" then obtain n :: int where n: "of_int n = vangle u v / pi" by (subst (asm) sin_zero_iff_int2) auto also have "\ \ {0..1}" using vangle_bounds by (auto simp: field_simps) finally have "n \ {0,1}" by auto thus "vangle u v \ {0,pi}" using n by (auto simp: field_simps) qed auto lemma sin_angle_zero_iff: "sin (angle a b c) = 0 \ angle a b c \ {0, pi}" unfolding angle_def by (simp only: sin_vangle_zero_iff) lemma vangle_collinear: "vangle u v \ {0, pi} \ collinear {0, u, v}" apply (subst norm_cauchy_schwarz_equal [symmetric]) apply (subst norm_cauchy_schwarz_abs_eq) apply (auto dest!: vangle_eq_0D vangle_eq_piD simp: eq_neg_iff_add_eq_0) done lemma angle_collinear: "angle a b c \ {0, pi} \ collinear {a, b, c}" apply (unfold angle_def, drule vangle_collinear) apply (subst collinear_translate_iff[symmetric, of _ "-b"]) apply (auto simp: insert_commute) done lemma not_collinear_vangle: "\collinear {0,u,v} \ vangle u v \ {0<.. vangle u v = pi") auto lemma not_collinear_angle: "\collinear {a,b,c} \ angle a b c \ {0<.. angle a b c = pi") auto subsection\Contributions from Lukas Bulwahn\ lemma vangle_scales: assumes "0 < c" shows "vangle (c *\<^sub>R v\<^sub>1) v\<^sub>2 = vangle v\<^sub>1 v\<^sub>2" using assms unfolding vangle_def by auto lemma vangle_inverse: "vangle (- v\<^sub>1) v\<^sub>2 = pi - vangle v\<^sub>1 v\<^sub>2" proof - have "\v\<^sub>1 \ v\<^sub>2 / (norm v\<^sub>1 * norm v\<^sub>2)\ \ 1" proof cases assume "v\<^sub>1 \ 0 \ v\<^sub>2 \ 0" from this show ?thesis by (simp add: Cauchy_Schwarz_ineq2) next assume "\ (v\<^sub>1 \ 0 \ v\<^sub>2 \ 0)" from this show ?thesis by auto qed from this show ?thesis unfolding vangle_def by (simp add: arccos_minus_abs) qed lemma orthogonal_iff_angle: shows "orthogonal (A - B) (C - B) \ angle A B C = pi / 2" unfolding angle_def by (auto simp only: orthogonal_iff_vangle) lemma angle_inverse: assumes "between (A, C) B" assumes "A \ B" "B \ C" shows "angle A B D = pi - angle C B D" proof - from \between (A, C) B\ obtain u where u: "u \ 0" "u \ 1" and X: "B = u *\<^sub>R A + (1 - u) *\<^sub>R C" by (metis add.commute betweenE between_commute) from \A \ B\ \B \ C\ X have "u \ 0" "u \ 1" by auto have "0 < ((1 - u) / u)" using \u \ 0\ \u \ 1\ \u \ 0\ \u \ 1\ by simp from X have "A - B = - (1 - u) *\<^sub>R (C - A)" by (simp add: real_vector.scale_right_diff_distrib real_vector.scale_left_diff_distrib) moreover from X have "C - B = u *\<^sub>R (C - A)" by (simp add: scaleR_diff_left real_vector.scale_right_diff_distrib) ultimately have "A - B = - (((1 - u) / u) *\<^sub>R (C - B))" using \u \ 0\ by simp (metis minus_diff_eq real_vector.scale_minus_left) from this have "vangle (A - B) (D - B) = pi - vangle (C - B) (D - B)" using \0 < (1 - u) / u\ by (simp add: vangle_inverse vangle_scales) from this show ?thesis unfolding angle_def by simp qed lemma strictly_between_implies_angle_eq_pi: assumes "between (A, C) B" assumes "A \ B" "B \ C" shows "angle A B C = pi" proof - from \between (A, C) B\ obtain u where u: "u \ 0" "u \ 1" and X: "B = u *\<^sub>R A + (1 - u) *\<^sub>R C" by (metis add.commute betweenE between_commute) from \A \ B\ \B \ C\ X have "u \ 0" "u \ 1" by auto from \A \ B\ \B \ C\ \between (A, C) B\ have "A \ C" by auto from X have "A - B = - (1 - u) *\<^sub>R (C - A)" by (simp add: real_vector.scale_right_diff_distrib real_vector.scale_left_diff_distrib) moreover from this have "dist A B = norm ((1 - u) *\<^sub>R (C - A))" using \u \ 0\ \u \ 1\ by (simp add: dist_norm) moreover from X have "C - B = u *\<^sub>R (C - A)" by (simp add: scaleR_diff_left real_vector.scale_right_diff_distrib) moreover from this have "dist C B = norm (u *\<^sub>R (C - A))" by (simp add: dist_norm) ultimately have "(A - B) \ (C - B) / (dist A B * dist C B) = u * (u - 1) / (\1 - u\ * \u\)" using \A \ C\ by (simp add: dot_square_norm power2_eq_square) also have "\ = - 1" using \u \ 0\ \u \ 1\ \u \ 0\ \u \ 1\ by (simp add: divide_eq_minus_1_iff) finally show ?thesis unfolding angle_altdef by simp qed end diff --git a/thys/UpDown_Scheme/Grid_Point.thy b/thys/UpDown_Scheme/Grid_Point.thy --- a/thys/UpDown_Scheme/Grid_Point.thy +++ b/thys/UpDown_Scheme/Grid_Point.thy @@ -1,186 +1,186 @@ section \ Grid Points \ theory Grid_Point -imports "HOL-Analysis.Ordered_Euclidean_Space" +imports "HOL-Analysis.Multivariate_Analysis" begin type_synonym grid_point = "(nat \ int) list" definition lv :: "grid_point \ nat \ nat" where "lv p d = fst (p ! d)" definition ix :: "grid_point \ nat \ int" where "ix p d = snd (p ! d)" definition level :: "grid_point \ nat" where "level p = (\ i < length p. lv p i)" lemma level_all_eq: assumes "\d. d < length p \ lv p d = lv p' d" and "length p = length p'" shows "level p' = level p" unfolding level_def using assms by auto datatype dir = left | right fun sgn :: "dir \ int" where "sgn left = -1" | "sgn right = 1" fun inv :: "dir \ dir" where "inv left = right" | "inv right = left" lemma inv_inv[simp]: "inv (inv dir) = dir" by (cases dir) simp_all lemma sgn_inv[simp]: "sgn (inv dir) = - sgn dir" by (cases dir, auto) definition child :: "grid_point \ dir \ nat \ grid_point" where "child p dir d = p[d := (lv p d + 1, 2 * (ix p d) + sgn dir)]" lemma child_length[simp]: "length (child p dir d) = length p" unfolding child_def by simp lemma child_odd[simp]: "d < length p \ odd (ix (child p dir d) d)" unfolding child_def ix_def by (cases dir, auto) lemma child_eq: "p ! d = (l, i) \ \ j. child p dir d = p[d := (l + 1, j)]" by (auto simp add: child_def ix_def lv_def) lemma child_other: "d \ d' \ child p dir d ! d' = p ! d'" unfolding child_def lv_def ix_def by (cases "d' < length p", auto) lemma child_invariant: assumes "d' < length p" shows "(child p dir d ! d' = p ! d') = (d \ d')" proof - obtain l i where "p ! d' = (l, i)" using prod.exhaust . with assms show ?thesis unfolding child_def ix_def lv_def by auto qed lemma child_single_level: "d < length p \ lv (child p dir d) d > lv p d" unfolding lv_def child_def by simp lemma child_lv: "d < length p \ lv (child p dir d) d = lv p d + 1" unfolding child_def lv_def by simp lemma child_lv_other: assumes "d' \ d" shows "lv (child p dir d') d = lv p d" using child_other[OF assms] unfolding lv_def by simp lemma child_ix_left: "d < length p \ ix (child p left d) d = 2 * ix p d - 1" unfolding child_def ix_def by simp lemma child_ix_right: "d < length p \ ix (child p right d) d = 2 * ix p d + 1" unfolding child_def ix_def by simp lemma child_ix: "d < length p \ ix (child p dir d) d = 2 * ix p d + sgn dir" unfolding child_def ix_def by simp lemma child_level[simp]: assumes "d < length p" shows "level (child p dir d) = level p + 1" proof - have inter: "{0.. {d} = {d}" using assms by auto have "level (child p dir d) = (\ d' = 0.. {d} then lv p d + 1 else lv p d')" by (auto intro!: sum.cong simp add: child_lv_other child_lv level_def) moreover have "level p + 1 = (\ d' = 0.. {d} then lv p d else lv p d') + 1" by (auto intro!: sum.cong simp add: child_lv_other child_lv level_def) ultimately show ?thesis unfolding sum.If_cases[OF finite_atLeastLessThan] inter using assms by auto qed lemma child_ex_neighbour: "\ b'. child b dir d = child b' (inv dir) d" proof show "child b dir d = child (b[d := (lv b d, ix b d + sgn dir)]) (inv dir) d" unfolding child_def ix_def lv_def by (cases "d < length b", auto simp add: algebra_simps) qed lemma child_level_gt[simp]: "level (child p dir d) >= level p" by (cases "d < length p", simp, simp add: child_def) lemma child_estimate_child: assumes "d < length p" and "l \ lv p d" and i'_range: "ix p d < (i + 1) * 2^(lv p d - l) \ ix p d > (i - 1) * 2^(lv p d - l)" (is "?top p \ ?bottom p") and is_child: "p' = child p dir d" shows "?top p' \ ?bottom p'" proof from is_child and \d < length p\ have "lv p' d = lv p d + 1" by (auto simp add: child_def ix_def lv_def) hence "lv p' d - l = lv p d - l + 1" using \lv p d >= l\ by auto hence pow_l'': "(2^(lv p' d - l) :: int) = 2 * 2^(lv p d - l)" by auto show "?top p'" proof - from is_child and \d < length p\ have "ix p' d \ 2 * (ix p d) + 1" by (cases dir, auto simp add: child_def lv_def ix_def) also have "\ < (i + 1) * (2 * 2^(lv p d - l))" using i'_range by auto finally show ?thesis using pow_l'' by auto qed show "?bottom p'" proof - have "(i - 1) * 2^(lv p' d - l) = 2 * ((i - 1) * 2^(lv p d - l))" using pow_l'' by simp also have "\ < 2 * ix p d - 1" using i'_range by auto finally show ?thesis using is_child and \d < length p\ by (cases dir, auto simp add: child_def lv_def ix_def) qed qed lemma child_neighbour: assumes "child p (inv dir) d = child ps dir d" (is "?c_p = ?c_ps") shows "ps = p[d := (lv p d, ix p d - sgn dir)]" (is "_ = ?ps") proof (rule nth_equalityI) have "length ?c_ps = length ?c_p" using \?c_p = ?c_ps\ by simp hence len_eq: "length ps = length p" by simp thus "length ps = length ?ps" by simp show "ps ! i = ?ps ! i" if "i < length ps" for i proof - have "i < length p" using that len_eq by auto show "ps ! i = ?ps ! i" proof (cases "d = i") case [simp]: True have "?c_p ! i = ?c_ps ! i" using \?c_p = ?c_ps\ by auto hence "ix p i = ix ps d + sgn dir" and "lv p i = lv ps i" by (auto simp add: child_def nth_list_update_eq[OF \i < length p\] nth_list_update_eq[OF \i < length ps\]) thus ?thesis by (simp add: lv_def ix_def \i < length p\) next assume "d \ i" with child_other[OF this, of ps dir] child_other[OF this, of p "inv dir"] show ?thesis using assms by auto qed qed qed definition start :: "nat \ grid_point" where "start dm = replicate dm (0, 1)" lemma start_lv[simp]: "d < dm \ lv (start dm) d = 0" unfolding start_def lv_def by simp lemma start_ix[simp]: "d < dm \ ix (start dm) d = 1" unfolding start_def ix_def by simp lemma start_length[simp]: "length (start dm) = dm" unfolding start_def by auto lemma level_start_0[simp]: "level (start dm) = 0" using level_def by auto end diff --git a/thys/Winding_Number_Eval/Cauchy_Index_Theorem.thy b/thys/Winding_Number_Eval/Cauchy_Index_Theorem.thy --- a/thys/Winding_Number_Eval/Cauchy_Index_Theorem.thy +++ b/thys/Winding_Number_Eval/Cauchy_Index_Theorem.thy @@ -1,4682 +1,4682 @@ (* Author: Wenda Li *) section \Cauchy's index theorem\ theory Cauchy_Index_Theorem imports - "HOL-Analysis.Ordered_Euclidean_Space" + "HOL-Analysis.Multivariate_Analysis" "Sturm_Tarski.Sturm_Tarski" "HOL-Computational_Algebra.Fundamental_Theorem_Algebra" Missing_Transcendental Missing_Algebraic Missing_Analysis begin text \This theory formalises Cauchy indices on the complex plane and relate them to winding numbers\ subsection \Misc\ (*refined version of the library one with the same name by dropping unnecessary assumptions*) lemma atMostAtLeast_subset_convex: fixes C :: "real set" assumes "convex C" and "x \ C" "y \ C" shows "{x .. y} \ C" proof safe fix z assume z: "z \ {x .. y}" have "z \ C" if *: "x < z" "z < y" proof - let ?\ = "(y - z) / (y - x)" have "0 \ ?\" "?\ \ 1" using assms * by (auto simp: field_simps) then have comb: "?\ * x + (1 - ?\) * y \ C" using assms iffD1[OF convex_alt, rule_format, of C y x ?\] by (simp add: algebra_simps) have "?\ * x + (1 - ?\) * y = (y - z) * x / (y - x) + (1 - (y - z) / (y - x)) * y" by (auto simp: field_simps) also have "\ = ((y - z) * x + (y - x - (y - z)) * y) / (y - x)" using * by (simp only: add_divide_distrib) (auto simp: field_simps) also have "\ = z" using assms * by (auto simp: field_simps) finally show ?thesis using comb by auto qed then show "z \ C" using z assms by (auto simp: le_less) qed lemma arg_elim: "f x \ x= y \ f y" by auto lemma arg_elim2: "f x1 x2 \ x1= y1 \x2=y2 \ f y1 y2" by auto lemma arg_elim3: "\f x1 x2 x3;x1= y1;x2=y2;x3=y3 \ \ f y1 y2 y3" by auto lemma IVT_strict: fixes f :: "'a::linear_continuum_topology \ 'b::linorder_topology" assumes "(f a > y \ y > f b) \ (f a < y \ y < f b)" "ax. a < x \ x < b \ f x = y" by (metis IVT' IVT2' assms(1) assms(2) assms(3) linorder_neq_iff order_le_less order_less_imp_le) lemma (in dense_linorder) atLeastAtMost_subseteq_greaterThanLessThan_iff: "{a .. b} \ { c <..< d } \ (a \ b \ c < a \ b < d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma Re_winding_number_half_right: assumes "\p\path_image \. Re p \ Re z" and "valid_path \" and "z\path_image \" shows "Re(winding_number \ z) = (Im (Ln (pathfinish \ - z)) - Im (Ln (pathstart \ - z)))/(2*pi)" proof - define g where "g=(\t. \ t - z)" define st fi where "st\pathstart g" and "fi\pathfinish g" have "valid_path g" "0\path_image g" and pos_img:"\p\path_image g. Re p \ 0" unfolding g_def subgoal using assms(2) by auto subgoal using assms(3) by auto subgoal using assms(1) by fastforce done have "(inverse has_contour_integral Ln fi - Ln st) g" unfolding fi_def st_def proof (rule contour_integral_primitive[OF _ \valid_path g\,of " - \\<^sub>\\<^sub>0"]) fix x::complex assume "x \ - \\<^sub>\\<^sub>0" then have "(Ln has_field_derivative inverse x) (at x)" using has_field_derivative_Ln by auto then show "(Ln has_field_derivative inverse x) (at x within - \\<^sub>\\<^sub>0)" using has_field_derivative_at_within by auto next show "path_image g \ - \\<^sub>\\<^sub>0" using pos_img \0\path_image g\ by (metis ComplI antisym assms(3) complex_nonpos_Reals_iff complex_surj subsetI zero_complex.code) qed then have winding_eq:"2*pi*\*winding_number g 0 = (Ln fi - Ln st)" using has_contour_integral_winding_number[OF \valid_path g\ \0\path_image g\ ,simplified,folded inverse_eq_divide] has_contour_integral_unique by auto have "Re(winding_number g 0) = (Im (Ln fi) - Im (Ln st))/(2*pi)" (is "?L=?R") proof - have "?L = Re((Ln fi - Ln st)/(2*pi*\))" using winding_eq[symmetric] by auto also have "... = ?R" by (metis Im_divide_of_real Im_i_times complex_i_not_zero minus_complex.simps(2) mult.commute mult_divide_mult_cancel_left_if times_divide_eq_right) finally show ?thesis . qed then show ?thesis unfolding g_def fi_def st_def using winding_number_offset by simp qed lemma Re_winding_number_half_upper: assumes pimage:"\p\path_image \. Im p \ Im z" and "valid_path \" and "z\path_image \" shows "Re(winding_number \ z) = (Im (Ln (\*z - \*pathfinish \)) - Im (Ln (\*z - \*pathstart \ )))/(2*pi)" proof - define \' where "\'=(\t. - \ * (\ t - z) + z)" have "Re (winding_number \' z) = (Im (Ln (pathfinish \' - z)) - Im (Ln (pathstart \' - z))) / (2 * pi)" unfolding \'_def apply (rule Re_winding_number_half_right) subgoal using pimage unfolding path_image_def by auto subgoal apply (rule valid_path_compose_holomorphic[OF \valid_path \\,of "\x. -\ * (x-z) + z" UNIV , unfolded comp_def]) by (auto intro!:holomorphic_intros) subgoal using \z\path_image \\ unfolding path_image_def by auto done moreover have "winding_number \' z = winding_number \ z" proof - define f where "f=(\x. -\ * (x-z) + z)" define c where "c= 1 / (complex_of_real (2 * pi) * \)" have "winding_number \' z = winding_number (f o \) z" unfolding \'_def comp_def f_def by auto also have "... = c * contour_integral \ (\w. deriv f w / (f w - z))" unfolding c_def proof (rule winding_number_comp[of UNIV]) show "z \ path_image (f \ \)" using \z\path_image \\ unfolding f_def path_image_def by auto qed (auto simp add:f_def \valid_path \\ intro!:holomorphic_intros) also have "... = c * contour_integral \ (\w. 1 / (w - z))" proof - have "deriv f x = -\" for x unfolding f_def by (auto intro!:derivative_eq_intros DERIV_imp_deriv) then show ?thesis unfolding f_def c_def by (auto simp add:field_simps divide_simps intro!:arg_cong2[where f=contour_integral]) qed also have "... = winding_number \ z" using winding_number_valid_path[OF \valid_path \\ \z\path_image \\,folded c_def] by simp finally show ?thesis . qed moreover have "pathfinish \' = z+ \*z -\* pathfinish \" "pathstart \' = z+ \*z -\*pathstart \" unfolding \'_def path_defs by (auto simp add:algebra_simps) ultimately show ?thesis by auto qed lemma Re_winding_number_half_lower: assumes pimage:"\p\path_image \. Im p \ Im z" and "valid_path \" and "z\path_image \" shows "Re(winding_number \ z) = (Im (Ln ( \*pathfinish \ - \*z)) - Im (Ln (\*pathstart \ - \*z)))/(2*pi)" proof - define \' where "\'=(\t. \ * (\ t - z) + z)" have "Re (winding_number \' z) = (Im (Ln (pathfinish \' - z)) - Im (Ln (pathstart \' - z))) / (2 * pi)" unfolding \'_def apply (rule Re_winding_number_half_right) subgoal using pimage unfolding path_image_def by auto subgoal apply (rule valid_path_compose_holomorphic[OF \valid_path \\,of "\x. \ * (x-z) + z" UNIV , unfolded comp_def]) by (auto intro!:holomorphic_intros) subgoal using \z\path_image \\ unfolding path_image_def by auto done moreover have "winding_number \' z = winding_number \ z" proof - define f where "f=(\x. \ * (x-z) + z)" define c where "c= 1 / (complex_of_real (2 * pi) * \)" have "winding_number \' z = winding_number (f o \) z" unfolding \'_def comp_def f_def by auto also have "... = c * contour_integral \ (\w. deriv f w / (f w - z))" unfolding c_def proof (rule winding_number_comp[of UNIV]) show "z \ path_image (f \ \)" using \z\path_image \\ unfolding f_def path_image_def by auto qed (auto simp add:f_def \valid_path \\ intro!:holomorphic_intros) also have "... = c * contour_integral \ (\w. 1 / (w - z))" proof - have "deriv f x = \" for x unfolding f_def by (auto intro!:derivative_eq_intros DERIV_imp_deriv) then show ?thesis unfolding f_def c_def by (auto simp add:field_simps divide_simps intro!:arg_cong2[where f=contour_integral]) qed also have "... = winding_number \ z" using winding_number_valid_path[OF \valid_path \\ \z\path_image \\,folded c_def] by simp finally show ?thesis . qed moreover have "pathfinish \' = z+ \* pathfinish \ - \*z" "pathstart \' = z+ \*pathstart \ - \*z" unfolding \'_def path_defs by (auto simp add:algebra_simps) ultimately show ?thesis by auto qed lemma Re_winding_number_half_left: assumes neg_img:"\p\path_image \. Re p \ Re z" and "valid_path \" and "z\path_image \" shows "Re(winding_number \ z) = (Im (Ln (z - pathfinish \)) - Im (Ln (z - pathstart \ )))/(2*pi)" proof - define \' where "\'\(\t. 2*z - \ t)" have "Re (winding_number \' z) = (Im (Ln (pathfinish \' - z)) - Im (Ln (pathstart \' - z))) / (2 * pi)" unfolding \'_def apply (rule Re_winding_number_half_right) subgoal using neg_img unfolding path_image_def by auto subgoal apply (rule valid_path_compose_holomorphic[OF \valid_path \\,of "\t. 2*z-t" UNIV, unfolded comp_def]) by (auto intro:holomorphic_intros) subgoal using \z\path_image \\ unfolding path_image_def by auto done moreover have "winding_number \' z = winding_number \ z" proof - define f where "f=(\t. 2*z-t)" define c where "c= 1 / (complex_of_real (2 * pi) * \)" have "winding_number \' z = winding_number (f o \) z" unfolding \'_def comp_def f_def by auto also have "... = c * contour_integral \ (\w. deriv f w / (f w - z))" unfolding c_def proof (rule winding_number_comp[of UNIV]) show "z \ path_image (f \ \)" using \z\path_image \\ unfolding f_def path_image_def by auto qed (auto simp add:f_def \valid_path \\ intro:holomorphic_intros) also have "... = c * contour_integral \ (\w. 1 / (w - z))" unfolding f_def c_def by (auto simp add:field_simps divide_simps intro!:arg_cong2[where f=contour_integral]) also have "... = winding_number \ z" using winding_number_valid_path[OF \valid_path \\ \z\path_image \\,folded c_def] by simp finally show ?thesis . qed moreover have "pathfinish \' = 2*z - pathfinish \" "pathstart \' = 2*z - pathstart \" unfolding \'_def path_defs by auto ultimately show ?thesis by auto qed lemma continuous_on_open_Collect_neq: fixes f g :: "'a::topological_space \ 'b::t2_space" assumes f: "continuous_on S f" and g: "continuous_on S g" and "open S" shows "open {x\S. f x \ g x}" proof (rule topological_space_class.openI) fix t assume "t \ {x\S. f x \ g x}" then obtain U0 V0 where "open U0" "open V0" "f t \ U0" "g t \ V0" "U0 \ V0 = {}" "t\S" by (auto simp add: separation_t2) obtain U1 where "open U1" "t \ U1" "\y\(S \ U1). f y \ U0" using f[unfolded continuous_on_topological,rule_format,OF \t\S\ \open U0\ \f t \U0\] by auto obtain V1 where "open V1" "t \ V1" "\y\(S \ V1). g y \ V0" using g[unfolded continuous_on_topological,rule_format,OF \t\S\ \open V0\ \g t \V0\] by auto define T where "T=V1 \ U1 \ S" have "open T" unfolding T_def using \open U1\ \open V1\ \open S\ by auto moreover have "t \ T" unfolding T_def using \t\U1\ \t\V1\ \t\S\ by auto moreover have "T \ {x \ S. f x \ g x}" unfolding T_def using \U0 \ V0 = {}\ \\y\S \ U1. f y \ U0\ \\y\S \ V1. g y \ V0\ by auto ultimately show "\T. open T \ t \ T \ T \ {x \ S. f x \ g x}" by auto qed subsection \Sign at a filter\ (*Relaxation in types could be done in the future.*) definition has_sgnx::"(real \ real) \ real \ real filter \ bool" (infixr "has'_sgnx" 55) where "(f has_sgnx c) F= (eventually (\x. sgn(f x) = c) F)" definition sgnx_able (infixr "sgnx'_able" 55) where "(f sgnx_able F) = (\c. (f has_sgnx c) F)" definition sgnx where "sgnx f F = (SOME c. (f has_sgnx c) F)" lemma has_sgnx_eq_rhs: "(f has_sgnx x) F \ x = y \ (f has_sgnx y) F" by simp named_theorems sgnx_intros "introduction rules for has_sgnx" setup \ Global_Theory.add_thms_dynamic (@{binding sgnx_eq_intros}, fn context => Named_Theorems.get (Context.proof_of context) @{named_theorems sgnx_intros} |> map_filter (try (fn thm => @{thm has_sgnx_eq_rhs} OF [thm]))) \ lemma sgnx_able_sgnx:"f sgnx_able F \ (f has_sgnx (sgnx f F)) F" unfolding sgnx_able_def sgnx_def using someI_ex by metis lemma has_sgnx_imp_sgnx_able[elim]: "(f has_sgnx c) F \ f sgnx_able F" unfolding sgnx_able_def by auto lemma has_sgnx_unique: assumes "F\bot" "(f has_sgnx c1) F" "(f has_sgnx c2) F" shows "c1=c2" proof (rule ccontr) assume "c1 \ c2 " have "eventually (\x. sgn(f x) = c1 \ sgn(f x) = c2) F" using assms unfolding has_sgnx_def eventually_conj_iff by simp then have "eventually (\_. c1 = c2) F" by (elim eventually_mono,auto) then have "eventually (\_. False) F" using \c1 \ c2\ by auto then show False using \F \ bot\ eventually_False by auto qed lemma has_sgnx_imp_sgnx[elim]: "(f has_sgnx c) F \F\bot \ sgnx f F = c" using has_sgnx_unique sgnx_def by auto lemma has_sgnx_const[simp,sgnx_intros]: "((\_. c) has_sgnx sgn c) F" by (simp add: has_sgnx_def) lemma finite_sgnx_at_left_at_right: assumes "finite {t. f t=0 \ a t{a<..0" "f sgnx_able (at_right x)" "sgnx f (at_right x)\0" proof - define ls where "ls \ {t. (f t=0 \ t\s) \ at(if ls = {} then (a+x)/2 else (Max ls + x)/2)" have "finite ls" proof - have "{t. f t=0 \ a t {t. f t=0 \ a t a ts \ a t a t {t. t\s \ a t a l a l{}" proof - have "Max ls \ ls" using assms(1,3) that \finite ls\ apply (intro linorder_class.Max_in) by auto then have "a Max ls < x" unfolding ls_def by auto then show ?thesis unfolding l_def using that x by auto qed ultimately show "l0" when t:"t\{l..a" using t \l>a\ by (meson atLeastLessThan_iff less_le_trans) then have "t\ls" using that t unfolding ls_def by auto then show False using True by auto qed then show ?thesis by auto next case False have "t>Max ls" using that False \l unfolding l_def by auto have False when "f t=0" proof - have "t>a" using t \l>a\ by (meson atLeastLessThan_iff less_le_trans) then have "t\ls" using that t unfolding ls_def by auto then have "t\Max ls" using \finite ls\ by auto then show False using \t>Max ls\ by auto qed then show ?thesis by auto qed have "(f has_sgnx sgn (f l)) (at_left x)" unfolding has_sgnx_def proof (rule eventually_at_leftI[OF _ \l]) fix t assume t:"t\{l<..a" "tl>a\ x by (meson greaterThanLessThan_iff less_trans)+ have False when "f t = 0" using noroot t that by auto moreover have False when "f l=0" using noroot t that by auto moreover have False when "f l>0 \ f t<0 \ f l <0 \ f t >0" proof - have False when "{l..t} \ s \{}" proof - obtain t' where t':"t'\{l..t}" "t'\s" using \{l..t} \ s \ {}\ by blast then have "a t'a < l\ atLeastAtMost_iff greaterThanLessThan_iff le_less less_trans t) then have "t'\ls" unfolding ls_def using \t'\s\ by auto then have "t'\Max ls" using \finite ls\ by auto moreover have "Max lsl \t'\ls\ \finite ls\ unfolding l_def by (auto simp add:ls_def ) ultimately show False using t'(1) by auto qed moreover have "{l..t} \ {a<..x>l. x < t \ f x = 0" apply (intro IVT_strict) using that t assms(2) by auto then obtain t' where "l{l..f t'=0\ by auto qed ultimately show "sgn (f t) = sgn (f l)" by (metis le_less not_less sgn_if) qed then show "f sgnx_able (at_left x)" by auto show "sgnx f (at_left x)\0" using noroot[of l,simplified] \(f has_sgnx sgn (f l)) (at_left x)\ by (simp add: has_sgnx_imp_sgnx sgn_if) next define rs where "rs \ {t. (f t=0 \ t\s) \ x t(if rs = {} then (x+b)/2 else (Min rs + x)/2)" have "finite rs" proof - have "{t. f t=0 \ x t {t. f t=0 \ a t x ts \ x t x t {t. t\s \ x tx" "ax \ a rx \ a r{}" proof - have "Min rs \ rs" using assms(1,3) that \finite rs\ apply (intro linorder_class.Min_in) by auto then have "x Min rs < b" unfolding rs_def by auto then show ?thesis unfolding r_def using that x by auto qed ultimately show "r>x" "a0" when t:"t\{x<..r}" for t proof (cases "rs = {}") case True have False when "f t=0" proof - have "tr using greaterThanAtMost_iff by fastforce then have "t\rs" using that t unfolding rs_def by auto then show False using True by auto qed then show ?thesis by auto next case False have "tr>x\ unfolding r_def by auto have False when "f t=0" proof - have "tr by (metis greaterThanAtMost_iff le_less less_trans) then have "t\rs" using that t unfolding rs_def by auto then have "t\Min rs" using \finite rs\ by auto then show False using \t by auto qed then show ?thesis by auto qed have "(f has_sgnx sgn (f r)) (at_right x)" unfolding has_sgnx_def proof (rule eventually_at_rightI[OF _ \r>x\]) fix t assume t:"t\{x<..a" "tr x by (meson greaterThanLessThan_iff less_trans)+ have False when "f t = 0" using noroot t that by auto moreover have False when "f r=0" using noroot t that by auto moreover have False when "f r>0 \ f t<0 \ f r <0 \ f t >0" proof - have False when "{t..r} \ s \{}" proof - obtain t' where t':"t'\{t..r}" "t'\s" using \{t..r} \ s \ {}\ by blast then have "x t'r < b\ atLeastAtMost_iff greaterThanLessThan_iff less_le_trans not_le t) then have "t'\rs" unfolding rs_def using t \t'\s\ by auto then have "t'\Min rs" using \finite rs\ by auto moreover have "Min rs>r" using \r>x\ \t'\rs\ \finite rs\ unfolding r_def by (auto simp add:rs_def ) ultimately show False using t'(1) by auto qed moreover have "{t..r} \ {a<..x>t. x < r \ f x = 0" apply (intro IVT_strict) using that t assms(2) by auto then obtain t' where "t{x<..r}" unfolding rs_def using t by auto then show False using noroot \f t'=0\ by auto qed ultimately show "sgn (f t) = sgn (f r)" by (metis le_less not_less sgn_if) qed then show "f sgnx_able (at_right x)" by auto show "sgnx f (at_right x)\0" using noroot[of r,simplified] \(f has_sgnx sgn (f r)) (at_right x)\ by (simp add: has_sgnx_imp_sgnx sgn_if) qed lemma sgnx_able_poly[simp]: "(poly p) sgnx_able (at_right a)" "(poly p) sgnx_able (at_left a)" "(poly p) sgnx_able at_top" "(poly p) sgnx_able at_bot" proof - show "(poly p) sgnx_able at_top" using has_sgnx_def poly_sgn_eventually_at_top sgnx_able_def by blast show "(poly p) sgnx_able at_bot" using has_sgnx_def poly_sgn_eventually_at_bot sgnx_able_def by blast show "(poly p) sgnx_able (at_right a)" proof (cases "p=0") case True then show ?thesis unfolding sgnx_able_def has_sgnx_def eventually_at_right using linordered_field_no_ub by force next case False obtain ub where "ub>a" and ub:"\z. az\ub\poly p z\0" using next_non_root_interval[OF False] by auto have "\z. az\ub\sgn(poly p z) = sgn (poly p ub)" proof (rule ccontr) assume "\ (\z. a < z \ z \ ub \ sgn (poly p z) = sgn (poly p ub))" then obtain z where "aub" "sgn(poly p z) \ sgn (poly p ub)" by auto moreover then have "poly p z\0" "poly p ub\0" "z\ub" using ub \ub>a\ by blast+ ultimately have "(poly p z>0 \ poly p ub<0) \ (poly p z<0 \ poly p ub>0)" by (metis linorder_neqE_linordered_idom sgn_neg sgn_pos) then have "\x>z. x < ub \ poly p x = 0" using poly_IVT_neg[of z ub p] poly_IVT_pos[of z ub p] \z\ub\ \z\ub\ by argo then show False using ub \a < z\ by auto qed then show ?thesis unfolding sgnx_able_def has_sgnx_def eventually_at_right apply (rule_tac exI[where x="sgn(poly p ub)"]) apply (rule_tac exI[where x="ub"]) using less_eq_real_def \ub>a\ by blast qed show "(poly p) sgnx_able (at_left a)" proof (cases "p=0") case True then show ?thesis unfolding sgnx_able_def has_sgnx_def eventually_at_right using linordered_field_no_ub by force next case False obtain lb where "lbz. lb\z\zpoly p z\0" using last_non_root_interval[OF False] by auto have "\z. lb\z\zsgn(poly p z) = sgn (poly p lb)" proof (rule ccontr) assume "\ (\z. lb\z\z sgn (poly p z) = sgn (poly p lb))" then obtain z where "lb\z" "z sgn (poly p lb)" by auto moreover then have "poly p z\0" "poly p lb\0" "z\lb" using ub \lb by blast+ ultimately have "(poly p z>0 \ poly p lb<0) \ (poly p z<0 \ poly p lb>0)" by (metis linorder_neqE_linordered_idom sgn_neg sgn_pos) then have "\x>lb. x < z \ poly p x = 0" using poly_IVT_neg[of lb z p] poly_IVT_pos[of lb z p] \lb\z\ \z\lb\ by argo then show False using ub \z < a\ by auto qed then show ?thesis unfolding sgnx_able_def has_sgnx_def eventually_at_left apply (rule_tac exI[where x="sgn(poly p lb)"]) apply (rule_tac exI[where x="lb"]) using less_eq_real_def \lb by blast qed qed lemma has_sgnx_identity[intro,sgnx_intros]: shows "x\0 \((\x. x) has_sgnx 1) (at_right x)" "x\0 \ ((\x. x) has_sgnx -1) (at_left x)" proof - show "x\0 \ ((\x. x) has_sgnx 1) (at_right x)" unfolding has_sgnx_def eventually_at_right apply (intro exI[where x="x+1"]) by auto show "x\0 \ ((\x. x) has_sgnx -1) (at_left x)" unfolding has_sgnx_def eventually_at_left apply (intro exI[where x="x-1"]) by auto qed lemma has_sgnx_divide[sgnx_intros]: assumes "(f has_sgnx c1) F" "(g has_sgnx c2) F" shows "((\x. f x / g x) has_sgnx c1 / c2) F" proof - have "\\<^sub>F x in F. sgn (f x) = c1 \ sgn (g x) = c2" using assms unfolding has_sgnx_def by (intro eventually_conj,auto) then have "\\<^sub>F x in F. sgn (f x / g x) = c1 / c2" apply (elim eventually_mono) by (simp add: sgn_mult sgn_divide) then show "((\x. f x / g x) has_sgnx c1 / c2) F" unfolding has_sgnx_def by auto qed lemma sgnx_able_divide[sgnx_intros]: assumes "f sgnx_able F" "g sgnx_able F" shows "(\x. f x / g x) sgnx_able F" using has_sgnx_divide by (meson assms(1) assms(2) sgnx_able_def) lemma sgnx_divide: assumes "F\bot" "f sgnx_able F" "g sgnx_able F" shows "sgnx (\x. f x / g x) F =sgnx f F / sgnx g F" proof - obtain c1 c2 where c1:"(f has_sgnx c1) F" and c2:"(g has_sgnx c2) F" using assms unfolding sgnx_able_def by auto have "sgnx f F=c1" "sgnx g F=c2" using c1 c2 \F\bot\ by auto moreover have "((\x. f x / g x) has_sgnx c1 / c2) F" using has_sgnx_divide[OF c1 c2] . ultimately show ?thesis using assms(1) has_sgnx_imp_sgnx by blast qed lemma has_sgnx_times[sgnx_intros]: assumes "(f has_sgnx c1) F" "(g has_sgnx c2) F" shows "((\x. f x* g x) has_sgnx c1 * c2) F" proof - have "\\<^sub>F x in F. sgn (f x) = c1 \ sgn (g x) = c2" using assms unfolding has_sgnx_def by (intro eventually_conj,auto) then have "\\<^sub>F x in F. sgn (f x * g x) = c1 * c2" apply (elim eventually_mono) by (simp add: sgn_mult) then show "((\x. f x* g x) has_sgnx c1 * c2) F" unfolding has_sgnx_def by auto qed lemma sgnx_able_times[sgnx_intros]: assumes "f sgnx_able F" "g sgnx_able F" shows "(\x. f x * g x) sgnx_able F" using has_sgnx_times by (meson assms(1) assms(2) sgnx_able_def) lemma sgnx_times: assumes "F\bot" "f sgnx_able F" "g sgnx_able F" shows "sgnx (\x. f x * g x) F =sgnx f F * sgnx g F" proof - obtain c1 c2 where c1:"(f has_sgnx c1) F" and c2:"(g has_sgnx c2) F" using assms unfolding sgnx_able_def by auto have "sgnx f F=c1" "sgnx g F=c2" using c1 c2 \F\bot\ by auto moreover have "((\x. f x* g x) has_sgnx c1 * c2) F" using has_sgnx_times[OF c1 c2] . ultimately show ?thesis using assms(1) has_sgnx_imp_sgnx by blast qed lemma tendsto_nonzero_has_sgnx: assumes "(f \ c) F" "c\0" shows "(f has_sgnx sgn c) F" proof (cases rule:linorder_cases[of c 0]) case less then have "\\<^sub>F x in F. f x<0" using order_topology_class.order_tendstoD[OF assms(1),of 0] by auto then show ?thesis unfolding has_sgnx_def apply (elim eventually_mono) using less by auto next case equal then show ?thesis using \c\0\ by auto next case greater then have "\\<^sub>F x in F. f x>0" using order_topology_class.order_tendstoD[OF assms(1),of 0] by auto then show ?thesis unfolding has_sgnx_def apply (elim eventually_mono) using greater by auto qed lemma tendsto_nonzero_sgnx: assumes "(f \ c) F" "F\bot" "c\0" shows "sgnx f F = sgn c" using tendsto_nonzero_has_sgnx by (simp add: assms has_sgnx_imp_sgnx) lemma filterlim_divide_at_bot_at_top_iff: assumes "(f \ c) F" "c\0" shows "(LIM x F. f x / g x :> at_bot) \ (g \ 0) F \ ((\x. g x) has_sgnx - sgn c) F" "(LIM x F. f x / g x :> at_top) \ (g \ 0) F \ ((\x. g x) has_sgnx sgn c) F" proof - show "(LIM x F. f x / g x :> at_bot) \ ((g \ 0) F ) \ ((\x. g x) has_sgnx - sgn c) F" proof assume asm:"LIM x F. f x / g x :> at_bot" then have "filterlim g (at 0) F" using filterlim_at_infinity_divide_iff[OF assms(1,2),of g] at_bot_le_at_infinity filterlim_mono by blast then have "(g \ 0) F" using filterlim_at by blast moreover have "(g has_sgnx - sgn c) F" proof - have "((\x. sgn c * inverse (f x)) \ sgn c * inverse c) F" using assms(1,2) by (auto intro:tendsto_intros) then have "LIM x F. sgn c * inverse (f x) * (f x / g x) :> at_bot" apply (elim filterlim_tendsto_pos_mult_at_bot[OF _ _ asm]) using \c\0\ sgn_real_def by auto then have "LIM x F. sgn c / g x :> at_bot" apply (elim filterlim_mono_eventually) using eventually_times_inverse_1[OF assms] by (auto elim:eventually_mono) then have "\\<^sub>F x in F. sgn c / g x < 0" using filterlim_at_bot_dense[of "\x. sgn c/g x" F] by auto then show ?thesis unfolding has_sgnx_def apply (elim eventually_mono) by (metis add.inverse_inverse divide_less_0_iff sgn_neg sgn_pos sgn_sgn) qed ultimately show "(g \ 0) F \ (g has_sgnx - sgn c) F" by auto next assume "(g \ 0) F \ (g has_sgnx - sgn c) F" then have asm:"(g \ 0) F" "(g has_sgnx - sgn c) F" by auto have "LIM x F. inverse (g x * sgn c) :> at_bot" proof (rule filterlim_inverse_at_bot) show "((\x. g x * sgn c) \ 0) F" apply (rule tendsto_mult_left_zero) using asm(1) by blast next show "\\<^sub>F x in F. g x * sgn c < 0" using asm(2) unfolding has_sgnx_def apply (elim eventually_mono) by (metis add.inverse_inverse assms(2) linorder_neqE_linordered_idom mult_less_0_iff neg_0_less_iff_less sgn_greater sgn_zero_iff) qed moreover have "((\x. f x * sgn c) \ c * sgn c) F" using \(f \ c) F\ \c\0\ apply (intro tendsto_intros) by (auto simp add:sgn_zero_iff) moreover have "c * sgn c >0" using \c\0\ by (simp add: sgn_real_def) ultimately have "LIM x F. (f x * sgn c) *inverse (g x * sgn c) :> at_bot" using filterlim_tendsto_pos_mult_at_bot by blast then show "LIM x F. f x / g x :> at_bot" using \c\0\ by (auto simp add:field_simps sgn_zero_iff) qed show "(LIM x F. f x / g x :> at_top) \ ((g \ 0) F ) \ ((\x. g x) has_sgnx sgn c) F" proof assume asm:"LIM x F. f x / g x :> at_top" then have "filterlim g (at 0) F" using filterlim_at_infinity_divide_iff[OF assms(1,2),of g] at_top_le_at_infinity filterlim_mono by blast then have "(g \ 0) F" using filterlim_at by blast moreover have "(g has_sgnx sgn c) F" proof - have "((\x. sgn c * inverse (f x)) \ sgn c * inverse c) F" using assms(1,2) by (auto intro:tendsto_intros) then have "LIM x F. sgn c * inverse (f x) * (f x / g x) :> at_top" apply (elim filterlim_tendsto_pos_mult_at_top[OF _ _ asm]) using \c\0\ sgn_real_def by auto then have "LIM x F. sgn c / g x :> at_top" apply (elim filterlim_mono_eventually) using eventually_times_inverse_1[OF assms] by (auto elim:eventually_mono) then have "\\<^sub>F x in F. sgn c / g x > 0" using filterlim_at_top_dense[of "\x. sgn c/g x" F] by auto then show ?thesis unfolding has_sgnx_def apply (elim eventually_mono) by (metis sgn_greater sgn_less sgn_neg sgn_pos zero_less_divide_iff) qed ultimately show "(g \ 0) F \ (g has_sgnx sgn c) F" by auto next assume "(g \ 0) F \ (g has_sgnx sgn c) F" then have asm:"(g \ 0) F" "(g has_sgnx sgn c) F" by auto have "LIM x F. inverse (g x * sgn c) :> at_top" proof (rule filterlim_inverse_at_top) show "((\x. g x * sgn c) \ 0) F" apply (rule tendsto_mult_left_zero) using asm(1) by blast next show "\\<^sub>F x in F. g x * sgn c > 0" using asm(2) unfolding has_sgnx_def apply (elim eventually_mono) by (metis assms(2) sgn_1_neg sgn_greater sgn_if zero_less_mult_iff) qed moreover have "((\x. f x * sgn c) \ c * sgn c) F" using \(f \ c) F\ \c\0\ apply (intro tendsto_intros) by (auto simp add:sgn_zero_iff) moreover have "c * sgn c >0" using \c\0\ by (simp add: sgn_real_def) ultimately have "LIM x F. (f x * sgn c) *inverse (g x * sgn c) :> at_top" using filterlim_tendsto_pos_mult_at_top by blast then show "LIM x F. f x / g x :> at_top" using \c\0\ by (auto simp add:field_simps sgn_zero_iff) qed qed lemma poly_sgnx_left_right: fixes c a::real and p::"real poly" assumes "p\0" shows "sgnx (poly p) (at_left a) = (if even (order a p) then sgnx (poly p) (at_right a) else -sgnx (poly p) (at_right a))" using assms proof (induction "degree p" arbitrary: p rule: less_induct) case less have ?case when "poly p a\0" proof - have "sgnx (poly p) (at_left a) = sgn (poly p a)" by (simp add: has_sgnx_imp_sgnx tendsto_nonzero_has_sgnx that) moreover have "sgnx (poly p) (at_right a) = sgn (poly p a)" by (simp add: has_sgnx_imp_sgnx tendsto_nonzero_has_sgnx that) moreover have "order a p = 0" using that by (simp add: order_0I) ultimately show ?thesis by auto qed moreover have ?case when "poly p a=0" proof - obtain q where pq:"p= [:-a,1:] * q" using \poly p a=0\ by (meson dvdE poly_eq_0_iff_dvd) then have "q\0" using \p\0\ by auto then have "degree q < degree p" unfolding pq by (subst degree_mult_eq,auto) have "sgnx (poly p) (at_left a) = - sgnx (poly q) (at_left a)" proof - have "sgnx (\x. poly p x) (at_left a) = sgnx (poly q) (at_left a) * sgnx (poly [:-a,1:]) (at_left a)" unfolding pq apply (subst poly_mult) apply (subst sgnx_times) by auto moreover have "sgnx (\x. poly [:-a,1:] x) (at_left a) = -1" apply (intro has_sgnx_imp_sgnx) unfolding has_sgnx_def eventually_at_left by (auto simp add: linordered_field_no_lb) ultimately show ?thesis by auto qed moreover have "sgnx (poly p) (at_right a) = sgnx (poly q) (at_right a)" proof - have "sgnx (\x. poly p x) (at_right a) = sgnx (poly q) (at_right a) * sgnx (poly [:-a,1:]) (at_right a)" unfolding pq apply (subst poly_mult) apply (subst sgnx_times) by auto moreover have "sgnx (\x. poly [:-a,1:] x) (at_right a) = 1" apply (intro has_sgnx_imp_sgnx) unfolding has_sgnx_def eventually_at_right by (auto simp add: linordered_field_no_ub) ultimately show ?thesis by auto qed moreover have "even (order a p) \ odd (order a q)" unfolding pq apply (subst order_mult[OF \p \ 0\[unfolded pq]]) using \q\0\ by (auto simp add:order_power_n_n[of _ 1, simplified]) moreover note less.hyps[OF \degree q < degree p\ \q\0\] ultimately show ?thesis by auto qed ultimately show ?case by blast qed lemma poly_has_sgnx_left_right: fixes c a::real and p::"real poly" assumes "p\0" shows "(poly p has_sgnx c) (at_left a) \ (if even (order a p) then (poly p has_sgnx c) (at_right a) else (poly p has_sgnx -c) (at_right a))" using poly_sgnx_left_right by (metis (no_types, hide_lams) add.inverse_inverse assms has_sgnx_unique sgnx_able_poly sgnx_able_sgnx trivial_limit_at_left_real trivial_limit_at_right_real) lemma sign_r_pos_sgnx_iff: "sign_r_pos p a \ sgnx (poly p) (at_right a) > 0" proof assume asm:"0 < sgnx (poly p) (at_right a)" obtain c where c_def:"(poly p has_sgnx c) (at_right a)" using sgnx_able_poly(1) sgnx_able_sgnx by blast then have "c>0" using asm using has_sgnx_imp_sgnx trivial_limit_at_right_real by blast then show "sign_r_pos p a" using c_def unfolding sign_r_pos_def has_sgnx_def apply (elim eventually_mono) by force next assume asm:"sign_r_pos p a" define c where "c = sgnx (poly p) (at_right a)" then have "(poly p has_sgnx c) (at_right a)" by (simp add: sgnx_able_sgnx) then have "(\\<^sub>F x in (at_right a). poly p x>0 \ sgn (poly p x) = c)" using asm unfolding has_sgnx_def sign_r_pos_def by (simp add:eventually_conj_iff) then have "\\<^sub>F x in (at_right a). c > 0" apply (elim eventually_mono) by fastforce then show "c>0" by auto qed lemma sgnx_values: assumes "f sgnx_able F" "F \ bot" shows "sgnx f F = -1 \ sgnx f F = 0 \ sgnx f F = 1" proof - obtain c where c_def:"(f has_sgnx c) F" using assms(1) unfolding sgnx_able_def by auto then obtain x where "sgn(f x) = c" unfolding has_sgnx_def using assms(2) eventually_happens by blast then have "c=-1 \ c=0 \ c=1" using sgn_if by metis moreover have "sgnx f F = c" using c_def by (simp add: assms(2) has_sgnx_imp_sgnx) ultimately show ?thesis by auto qed lemma has_sgnx_poly_at_top: "(poly p has_sgnx sgn_pos_inf p) at_top" using has_sgnx_def poly_sgn_eventually_at_top by blast lemma has_sgnx_poly_at_bot: "(poly p has_sgnx sgn_neg_inf p) at_bot" using has_sgnx_def poly_sgn_eventually_at_bot by blast lemma sgnx_poly_at_top: "sgnx (poly p) at_top = sgn_pos_inf p" by (simp add: has_sgnx_def has_sgnx_imp_sgnx poly_sgn_eventually_at_top) lemma sgnx_poly_at_bot: "sgnx (poly p) at_bot = sgn_neg_inf p" by (simp add: has_sgnx_def has_sgnx_imp_sgnx poly_sgn_eventually_at_bot) lemma poly_has_sgnx_values: assumes "p\0" shows "(poly p has_sgnx 1) (at_left a) \ (poly p has_sgnx - 1) (at_left a)" "(poly p has_sgnx 1) (at_right a) \ (poly p has_sgnx - 1) (at_right a)" "(poly p has_sgnx 1) at_top \ (poly p has_sgnx - 1) at_top" "(poly p has_sgnx 1) at_bot \ (poly p has_sgnx - 1) at_bot" proof - have "sgn_pos_inf p = 1 \ sgn_pos_inf p = -1" unfolding sgn_pos_inf_def by (simp add: assms sgn_if) then show "(poly p has_sgnx 1) at_top \ (poly p has_sgnx - 1) at_top" using has_sgnx_poly_at_top by metis next have "sgn_neg_inf p = 1 \ sgn_neg_inf p = -1" unfolding sgn_neg_inf_def by (simp add: assms sgn_if) then show "(poly p has_sgnx 1) at_bot \ (poly p has_sgnx - 1) at_bot" using has_sgnx_poly_at_bot by metis next obtain c where c_def:"(poly p has_sgnx c) (at_left a)" using sgnx_able_poly(2) sgnx_able_sgnx by blast then have "sgnx (poly p) (at_left a) = c" using assms by auto then have "c=-1 \ c=0 \ c=1" using sgnx_values sgnx_able_poly(2) trivial_limit_at_left_real by blast moreover have False when "c=0" proof - have "(poly p has_sgnx 0) (at_left a)" using c_def that by auto then obtain lb where "lby. (lb y < a) \ poly p y = 0" unfolding has_sgnx_def eventually_at_left sgn_if by (metis one_neq_zero zero_neq_neg_one) then have "{lb<.. proots p" unfolding proots_within_def by auto then have "infinite (proots p)" apply (elim infinite_super) using \lb by auto moreover have "finite (proots p)" using finite_proots[OF \p\0\] by auto ultimately show False by auto qed ultimately have "c=-1 \ c=1" by auto then show "(poly p has_sgnx 1) (at_left a) \ (poly p has_sgnx - 1) (at_left a)" using c_def by auto next obtain c where c_def:"(poly p has_sgnx c) (at_right a)" using sgnx_able_poly(1) sgnx_able_sgnx by blast then have "sgnx (poly p) (at_right a) = c" using assms by auto then have "c=-1 \ c=0 \ c=1" using sgnx_values sgnx_able_poly(1) trivial_limit_at_right_real by blast moreover have False when "c=0" proof - have "(poly p has_sgnx 0) (at_right a)" using c_def that by auto then obtain ub where "ub>a" "\y. (a y < ub) \ poly p y = 0" unfolding has_sgnx_def eventually_at_right sgn_if by (metis one_neq_zero zero_neq_neg_one) then have "{a<.. proots p" unfolding proots_within_def by auto then have "infinite (proots p)" apply (elim infinite_super) using \ub>a\ by auto moreover have "finite (proots p)" using finite_proots[OF \p\0\] by auto ultimately show False by auto qed ultimately have "c=-1 \ c=1" by auto then show "(poly p has_sgnx 1) (at_right a) \ (poly p has_sgnx - 1) (at_right a)" using c_def by auto qed lemma poly_sgnx_values: assumes "p\0" shows "sgnx (poly p) (at_left a) = 1 \ sgnx (poly p) (at_left a) = -1" "sgnx (poly p) (at_right a) = 1 \ sgnx (poly p) (at_right a) = -1" using poly_has_sgnx_values[OF \p\0\] has_sgnx_imp_sgnx trivial_limit_at_left_real trivial_limit_at_right_real by blast+ lemma has_sgnx_inverse: "(f has_sgnx c) F \ ((inverse o f) has_sgnx (inverse c)) F" unfolding has_sgnx_def comp_def apply (rule eventually_subst) apply (rule always_eventually) by (metis inverse_inverse_eq sgn_inverse) lemma has_sgnx_derivative_at_left: assumes g_deriv:"(g has_field_derivative c) (at x)" and "g x=0" and "c\0" shows "(g has_sgnx - sgn c) (at_left x)" proof - have "(g has_sgnx -1) (at_left x)" when "c>0" proof - obtain d1 where "d1>0" and d1_def:"\h>0. h < d1 \ g (x - h) < g x" using DERIV_pos_inc_left[OF g_deriv \c>0\] \g x=0\ by auto have "(g has_sgnx -1) (at_left x)" unfolding has_sgnx_def eventually_at_left apply (intro exI[where x="x-d1"]) using \d1>0\ d1_def by (metis (no_types, hide_lams) add.commute add_uminus_conv_diff assms(2) diff_add_cancel diff_strict_left_mono diff_zero minus_diff_eq sgn_neg) thus ?thesis by auto qed moreover have "(g has_sgnx 1) (at_left x)" when "c<0" proof - obtain d1 where "d1>0" and d1_def:"\h>0. h < d1 \ g (x - h) > g x" using DERIV_neg_dec_left[OF g_deriv \c<0\] \g x=0\ by auto have "(g has_sgnx 1) (at_left x)" unfolding has_sgnx_def eventually_at_left apply (intro exI[where x="x-d1"]) using \d1>0\ d1_def by (metis (no_types, hide_lams) add.commute add_uminus_conv_diff assms(2) diff_add_cancel diff_zero less_diff_eq minus_diff_eq sgn_pos) thus ?thesis using \c<0\ by auto qed ultimately show ?thesis using \c\0\ using sgn_real_def by auto qed lemma has_sgnx_derivative_at_right: assumes g_deriv:"(g has_field_derivative c) (at x)" and "g x=0" and "c\0" shows "(g has_sgnx sgn c) (at_right x)" proof - have "(g has_sgnx 1) (at_right x)" when "c>0" proof - obtain d2 where "d2>0" and d2_def:"\h>0. h < d2 \ g x < g (x + h)" using DERIV_pos_inc_right[OF g_deriv \c>0\] \g x=0\ by auto have "(g has_sgnx 1) (at_right x)" unfolding has_sgnx_def eventually_at_right apply (intro exI[where x="x+d2"]) using \d2>0\ d2_def by (metis add.commute assms(2) diff_add_cancel diff_less_eq less_add_same_cancel1 sgn_pos) thus ?thesis using \c>0\ by auto qed moreover have "(g has_sgnx -1) (at_right x)" when "c<0" proof - obtain d2 where "d2>0" and d2_def:"\h>0. h < d2 \ g x > g (x + h)" using DERIV_neg_dec_right[OF g_deriv \c<0\] \g x=0\ by auto have "(g has_sgnx -1) (at_right x)" unfolding has_sgnx_def eventually_at_right apply (intro exI[where x="x+d2"]) using \d2>0\ d2_def by (metis (no_types, hide_lams) add.commute add.right_inverse add_uminus_conv_diff assms(2) diff_add_cancel diff_less_eq sgn_neg) thus ?thesis using \c<0\ by auto qed ultimately show ?thesis using \c\0\ using sgn_real_def by auto qed lemma has_sgnx_split: "(f has_sgnx c) (at x) \ (f has_sgnx c) (at_left x) \ (f has_sgnx c) (at_right x)" unfolding has_sgnx_def using eventually_at_split by auto lemma sgnx_at_top_IVT: assumes "sgnx (poly p) (at_right a) \ sgnx (poly p) at_top" shows "\x>a. poly p x=0" proof (cases "p=0") case True then show ?thesis using gt_ex[of a] by simp next case False from poly_has_sgnx_values[OF this] have "(poly p has_sgnx 1) (at_right a) \ (poly p has_sgnx - 1) (at_right a)" "(poly p has_sgnx 1) at_top \ (poly p has_sgnx - 1) at_top" by auto moreover have ?thesis when has_r:"(poly p has_sgnx 1) (at_right a)" and has_top:"(poly p has_sgnx -1) at_top" proof - obtain b where "b>a" "poly p b>0" proof - obtain a' where "a'>a" and a'_def:"\y>a. y < a' \ sgn (poly p y) = 1" using has_r[unfolded has_sgnx_def eventually_at_right] by auto define b where "b=(a+a')/2" have "aa'>a\ by auto moreover have "poly p b>0" using a'_def[rule_format,OF \b>a\ \b] unfolding sgn_if by argo ultimately show ?thesis using that by auto qed moreover obtain c where "c>b" "poly p c<0" proof - obtain b' where b'_def:"\n\b'. sgn (poly p n) = - 1" using has_top[unfolded has_sgnx_def eventually_at_top_linorder] by auto define c where "c=1+max b b'" have "c>b" "c\b'" unfolding c_def using \b>a\ by auto moreover have "poly p c<0" using b'_def[rule_format,OF \b'\c\] unfolding sgn_if by argo ultimately show ?thesis using that by auto qed ultimately show ?thesis using poly_IVT_neg[of b c p] not_less by fastforce qed moreover have ?thesis when has_r:"(poly p has_sgnx -1) (at_right a)" and has_top:"(poly p has_sgnx 1) at_top" proof - obtain b where "b>a" "poly p b<0" proof - obtain a' where "a'>a" and a'_def:"\y>a. y < a' \ sgn (poly p y) = -1" using has_r[unfolded has_sgnx_def eventually_at_right] by auto define b where "b=(a+a')/2" have "aa'>a\ by auto moreover have "poly p b<0" using a'_def[rule_format,OF \b>a\ \b] unfolding sgn_if by argo ultimately show ?thesis using that by auto qed moreover obtain c where "c>b" "poly p c>0" proof - obtain b' where b'_def:"\n\b'. sgn (poly p n) = 1" using has_top[unfolded has_sgnx_def eventually_at_top_linorder] by auto define c where "c=1+max b b'" have "c>b" "c\b'" unfolding c_def using \b>a\ by auto moreover have "poly p c>0" using b'_def[rule_format,OF \b'\c\] unfolding sgn_if by argo ultimately show ?thesis using that by auto qed ultimately show ?thesis using poly_IVT_pos[of b c p] not_less by fastforce qed moreover have ?thesis when "(poly p has_sgnx 1) (at_right a) \ (poly p has_sgnx 1) at_top \ (poly p has_sgnx - 1) (at_right a) \ (poly p has_sgnx -1) at_top" proof - have "sgnx (poly p) (at_right a) = sgnx (poly p) at_top" using that has_sgnx_imp_sgnx by auto then have False using assms by simp then show ?thesis by auto qed ultimately show ?thesis by blast qed lemma sgnx_at_left_at_right_IVT: assumes "sgnx (poly p) (at_right a) \ sgnx (poly p) (at_left b)" "ax. a x poly p x=0" proof (cases "p=0") case True then show ?thesis using \a by (auto intro:exI[where x="(a+b)/2"]) next case False from poly_has_sgnx_values[OF this] have "(poly p has_sgnx 1) (at_right a) \ (poly p has_sgnx - 1) (at_right a)" "(poly p has_sgnx 1) (at_left b) \ (poly p has_sgnx - 1) (at_left b)" by auto moreover have ?thesis when has_r:"(poly p has_sgnx 1) (at_right a)" and has_l:"(poly p has_sgnx -1) (at_left b)" proof - obtain c where "a0" proof - obtain a' where "a'>a" and a'_def:"\y>a. y < a' \ sgn (poly p y) = 1" using has_r[unfolded has_sgnx_def eventually_at_right] by auto define c where "c=(a+min a' b)/2" have "aa'>a\ \b>a\ by auto moreover have "poly p c>0" using a'_def[rule_format,OF \c>a\ \c] unfolding sgn_if by argo ultimately show ?thesis using that by auto qed moreover obtain d where "cy>b'. y < b \ sgn (poly p y) = - 1" using has_l[unfolded has_sgnx_def eventually_at_left] by auto define d where "d=(b+max b' c)/2" have "b'c" unfolding d_def using \b>b'\ \b>c\ by auto moreover have "poly p d<0" using b'_def[rule_format, OF \b' \d] unfolding sgn_if by argo ultimately show ?thesis using that by auto qed ultimately obtain x where "cc>a\ \d by (auto intro: exI[where x=x]) qed moreover have ?thesis when has_r:"(poly p has_sgnx -1) (at_right a)" and has_l:"(poly p has_sgnx 1) (at_left b)" proof - obtain c where "aa" and a'_def:"\y>a. y < a' \ sgn (poly p y) = -1" using has_r[unfolded has_sgnx_def eventually_at_right] by auto define c where "c=(a+min a' b)/2" have "aa'>a\ \b>a\ by auto moreover have "poly p c<0" using a'_def[rule_format,OF \c>a\ \c] unfolding sgn_if by argo ultimately show ?thesis using that by auto qed moreover obtain d where "c0" proof - obtain b' where "b'y>b'. y < b \ sgn (poly p y) = 1" using has_l[unfolded has_sgnx_def eventually_at_left] by auto define d where "d=(b+max b' c)/2" have "b'c" unfolding d_def using \b>b'\ \b>c\ by auto moreover have "poly p d>0" using b'_def[rule_format, OF \b' \d] unfolding sgn_if by argo ultimately show ?thesis using that by auto qed ultimately obtain x where "cc>a\ \d by (auto intro: exI[where x=x]) qed moreover have ?thesis when "(poly p has_sgnx 1) (at_right a) \ (poly p has_sgnx 1) (at_left b) \ (poly p has_sgnx - 1) (at_right a) \ (poly p has_sgnx -1) (at_left b)" proof - have "sgnx (poly p) (at_right a) = sgnx (poly p) (at_left b)" using that has_sgnx_imp_sgnx by auto then have False using assms by simp then show ?thesis by auto qed ultimately show ?thesis by blast qed lemma sgnx_at_bot_IVT: assumes "sgnx (poly p) (at_left a) \ sgnx (poly p) at_bot" shows "\x (poly p has_sgnx - 1) (at_left a)" "(poly p has_sgnx 1) at_bot \ (poly p has_sgnx - 1) at_bot" by auto moreover have ?thesis when has_l:"(poly p has_sgnx 1) (at_left a)" and has_bot:"(poly p has_sgnx -1) at_bot" proof - obtain b where "b0" proof - obtain a' where "a'y>a'. y < a \ sgn (poly p y) = 1" using has_l[unfolded has_sgnx_def eventually_at_left] by auto define b where "b=(a+a')/2" have "a>b" "b>a'" unfolding b_def using \a' by auto moreover have "poly p b>0" using a'_def[rule_format,OF \b>a'\ \b] unfolding sgn_if by argo ultimately show ?thesis using that by auto qed moreover obtain c where "cn\b'. sgn (poly p n) = - 1" using has_bot[unfolded has_sgnx_def eventually_at_bot_linorder] by auto define c where "c=min b b'- 1" have "cb'" unfolding c_def using \b by auto moreover have "poly p c<0" using b'_def[rule_format,OF \b'\c\] unfolding sgn_if by argo ultimately show ?thesis using that by auto qed ultimately show ?thesis using poly_IVT_pos[of c b p] using not_less by fastforce qed moreover have ?thesis when has_l:"(poly p has_sgnx -1) (at_left a)" and has_bot:"(poly p has_sgnx 1) at_bot" proof - obtain b where "by>a'. y < a \ sgn (poly p y) = -1" using has_l[unfolded has_sgnx_def eventually_at_left] by auto define b where "b=(a+a')/2" have "a>b" "b>a'" unfolding b_def using \a' by auto moreover have "poly p b<0" using a'_def[rule_format,OF \b>a'\ \b] unfolding sgn_if by argo ultimately show ?thesis using that by auto qed moreover obtain c where "c0" proof - obtain b' where b'_def:"\n\b'. sgn (poly p n) = 1" using has_bot[unfolded has_sgnx_def eventually_at_bot_linorder] by auto define c where "c=min b b'- 1" have "cb'" unfolding c_def using \b by auto moreover have "poly p c>0" using b'_def[rule_format,OF \b'\c\] unfolding sgn_if by argo ultimately show ?thesis using that by auto qed ultimately show ?thesis using poly_IVT_neg[of c b p] using not_less by fastforce qed moreover have ?thesis when "(poly p has_sgnx 1) (at_left a) \ (poly p has_sgnx 1) at_bot \ (poly p has_sgnx - 1) (at_left a) \ (poly p has_sgnx -1) at_bot" proof - have "sgnx (poly p) (at_left a) = sgnx (poly p) at_bot" using that has_sgnx_imp_sgnx by auto then have False using assms by simp then show ?thesis by auto qed ultimately show ?thesis by blast qed lemma sgnx_poly_nz: assumes "poly p x\0" shows "sgnx (poly p) (at_left x) = sgn (poly p x)" "sgnx (poly p) (at_right x) = sgn (poly p x)" proof - have "(poly p has_sgnx sgn(poly p x)) (at x)" apply (rule tendsto_nonzero_has_sgnx) using assms by auto then show "sgnx (poly p) (at_left x) = sgn (poly p x)" "sgnx (poly p) (at_right x) = sgn (poly p x)" unfolding has_sgnx_split by auto qed subsection \Finite predicate segments over an interval\ inductive finite_Psegments::"(real \ bool) \ real \ real \ bool" for P where emptyI: "a\b \ finite_Psegments P a b"| insertI_1: "\s\{a..P s;\t\{s<.. \ finite_Psegments P a b"| insertI_2: "\s\{a..P s;(\t\{s<..P t);finite_Psegments P a s\ \ finite_Psegments P a b" lemma finite_Psegments_pos_linear: assumes "finite_Psegments P (b*lb+c) (b*ub+c) " and "b>0" shows "finite_Psegments (P o (\t. b*t+c)) lb ub" proof - have [simp]:"b\0" using \b>0\ by auto show ?thesis proof (rule finite_Psegments.induct[OF assms(1), of "\lb' ub'. finite_Psegments (P o (\t. b*t+c)) ((lb'-c)/b) ((ub'-c)/b)",simplified]) (*really weird application of the induction rule, is there an alternative?*) fix lb ub f assume "(lb::real)\ub" then have "(lb - c) / b \ (ub - c) / b" using \b>0\ by (auto simp add:field_simps) then show "finite_Psegments (f \ (\t. b * t + c)) ((ub - c) / b) ((lb - c) / b)" by (rule finite_Psegments.emptyI) next fix s lb ub P assume asm: "lb \ s \ s < ub" "\t\{s<.. (\t. b * t + c)) ((lb - c) / b) ((s - c) / b)" "s = lb \ P s" show "finite_Psegments (P \ (\t. b * t + c)) ((lb - c) / b) ((ub - c) / b)" apply (rule finite_Psegments.insertI_1[of "(s-c)/b"]) using asm \b>0\ by (auto simp add:field_simps) next fix s lb ub P assume asm: "lb \ s \ s < ub" "\t\{s<.. P t" "finite_Psegments (P \ (\t. b * t + c)) ((lb - c) / b) ((s - c) / b)" "s=lb \ P s" show "finite_Psegments (P \ (\t. b * t + c)) ((lb - c) / b) ((ub - c) / b)" apply (rule finite_Psegments.insertI_2[of "(s-c)/b"]) using asm \b>0\ by (auto simp add:field_simps) qed qed lemma finite_Psegments_congE: assumes "finite_Psegments Q lb ub" "\t. \lb \ Q t \ P t " shows "finite_Psegments P lb ub" using assms proof (induct rule:finite_Psegments.induct) case (emptyI a b) then show ?case using finite_Psegments.emptyI by auto next case (insertI_1 s a b) show ?case proof (rule finite_Psegments.insertI_1[of s]) have "P s" when "s\a" proof - have "s\{a<..s \ {a.. that by auto then show ?thesis using insertI_1 by auto qed then show "s = a \ P s" by auto next show "s \ {a..t\{s<..a" proof - have "s\{a<..s \ {a.. that by auto then show ?thesis using insertI_2 by auto qed then show "s = a \ P s" by auto next show "s \ {a..t\{s<.. P t" "finite_Psegments P a s" using insertI_2 by auto qed qed lemma finite_Psegments_constI: assumes "\t. \a \ P t = c" shows "finite_Psegments P a b" proof - have "finite_Psegments (\_. c) a b" proof - have ?thesis when "a\b" using that finite_Psegments.emptyI by auto moreover have ?thesis when "ac" apply (rule finite_Psegments.insertI_2[of a]) using that by (auto intro: finite_Psegments.emptyI) ultimately show ?thesis by argo qed then show ?thesis apply (elim finite_Psegments_congE) using assms by auto qed context begin private lemma finite_Psegments_less_eq1: assumes "finite_Psegments P a c" "b\c" shows "finite_Psegments P a b" using assms proof (induct arbitrary: b rule:finite_Psegments.induct) case (emptyI a c) then show ?case using finite_Psegments.emptyI by auto next case (insertI_1 s a c) have ?case when "b\s" using insertI_1 that by auto moreover have ?case when "b>s" proof - have "s \ {a..s \ {a.. \b \ c\ by auto moreover have "\t\{s<..\t\{s<.. that \b \ c\ by auto ultimately show ?case using finite_Psegments.insertI_1[OF _ _ _ \finite_Psegments P a s\] \ s = a \ P s\ by auto qed ultimately show ?case by fastforce next case (insertI_2 s a c) have ?case when "b\s" using insertI_2 that by auto moreover have ?case when "b>s" proof - have "s \ {a..s \ {a.. \b \ c\ by auto moreover have "\t\{s<.. P t" using \\t\{s<.. P t\ that \b \ c\ by auto ultimately show ?case using finite_Psegments.insertI_2[OF _ _ _ \finite_Psegments P a s\] \ s = a \ P s\ by auto qed ultimately show ?case by fastforce qed private lemma finite_Psegments_less_eq2: assumes "finite_Psegments P a c" "a\b" shows "finite_Psegments P b c" using assms proof (induct arbitrary: rule:finite_Psegments.induct) case (emptyI a c) then show ?case using finite_Psegments.emptyI by auto next case (insertI_1 s a c) have ?case when "s\b" proof - have "\t\{b<..b" apply (rule finite_Psegments.insertI_1[where s=s]) using insertI_1 that by auto ultimately show ?case by linarith next case (insertI_2 s a c) have ?case when "s\b" proof - have "\t\{b<.. P t" using insertI_2 that by auto then show ?thesis by (metis finite_Psegments_constI greaterThanLessThan_iff) qed moreover have ?case when "s>b" apply (rule finite_Psegments.insertI_2[where s=s]) using insertI_2 that by auto ultimately show ?case by linarith qed lemma finite_Psegments_included: assumes "finite_Psegments P a d" "a\b" "c\d" shows "finite_Psegments P b c" using finite_Psegments_less_eq2 finite_Psegments_less_eq1 assms by blast end lemma finite_Psegments_combine: assumes "finite_Psegments P a b" "finite_Psegments P b c" "b\{a..c}" "closed ({x. P x} \ {a..c})" shows "finite_Psegments P a c" using assms(2,1,3,4) proof (induct rule:finite_Psegments.induct) case (emptyI b c) then show ?case using finite_Psegments_included by auto next case (insertI_1 s b c) have "P s" proof - have "s {s..(s+c)/2}" have "closed S" proof - have "closed ({a. P a} \ {a..c})" using insertI_1(8) . moreover have "S = ({a. P a} \ {a..c}) \ {s..(s+c)/2}" using insertI_1(1,7) unfolding S_def by (auto simp add:field_simps) ultimately show ?thesis using closed_Int[of "{a. P a} \ {a..c}" "{s..(s+c)/2}"] by blast qed moreover have "\y\S. dist y s < e" when "e>0" for e proof - define y where "y = min ((s+c)/2) (e/2+s)" have "y\S" proof - have "y\{s..(s+c)/2}" unfolding y_def using \e>0\ \s by (auto simp add:min_mult_distrib_left algebra_simps) moreover have "P y" apply (rule insertI_1(3)[rule_format]) unfolding y_def using \e>0\ \s by (auto simp add:algebra_simps min_mult_distrib_left min_less_iff_disj) ultimately show ?thesis unfolding S_def by auto qed moreover have "dist y s e>0\ \s by (auto simp add:algebra_simps min_mult_distrib_left min_less_iff_disj dist_real_def) ultimately show ?thesis by auto qed ultimately have "s\S" using closed_approachable by auto then show ?thesis unfolding S_def by auto qed show ?case proof (rule finite_Psegments.insertI_1[of s]) show " s \ {a.. P s" "\t\{s<..P s\ by auto next have "closed ({a. P a} \ {a..s})" using closed_Int[OF \closed ({a. P a} \ {a..c})\,of "{a..s}",simplified] apply (elim arg_elim[of closed]) using \s \ {b.. \b \ {a..c}\ by auto then show "finite_Psegments P a s" using insertI_1 by auto qed next case (insertI_2 s b c) have ?case when "P s" proof (rule finite_Psegments.insertI_2[of s]) show "s \ {a.. P s" "\t\{s<.. P t" using that insertI_2 by auto next have "closed ({a. P a} \ {a..s})" using closed_Int[OF \closed ({a. P a} \ {a..c})\,of "{a..s}",simplified] apply (elim arg_elim[of closed]) using \s \ {b.. \b \ {a..c}\ by auto then show "finite_Psegments P a s" using insertI_2 by auto qed moreover have ?case when "\ P s" "s=b" using \finite_Psegments P a b\ proof (cases rule:finite_Psegments.cases) case emptyI then show ?thesis using insertI_2 that by (metis antisym_conv atLeastAtMost_iff finite_Psegments.insertI_2) next case (insertI_1 s0) have "P s" proof - have "s0 {(s0+s)/2..s}" have "closed S" using closed_Int[OF \closed ({a. P a} \ {a..c})\,of "{(s0+s)/2..s}",simplified] apply (elim arg_elim[of closed]) unfolding S_def using \s0 \ {a.. \ s \ {b.. \b \ {a..c}\ by auto moreover have "\y\S. dist y s < e" when "e>0" for e proof - define y where "y = max ((s+s0)/2) (s-e/2)" have "y\S" proof - have "y\{(s0+s)/2..s}" unfolding y_def using \e>0\ \s0 by (auto simp add:field_simps min_mult_distrib_left) moreover have "P y" apply (rule insertI_1(3)[rule_format]) unfolding y_def using \e>0\ \s0 \s=b\ by (auto simp add:field_simps max_mult_distrib_left less_max_iff_disj) ultimately show ?thesis unfolding S_def by auto qed moreover have "dist y s e>0\ \s0 by (auto simp add:algebra_simps max_mult_distrib_left less_max_iff_disj dist_real_def max_add_distrib_right) ultimately show ?thesis by auto qed ultimately have "s\S" using closed_approachable by auto then show ?thesis unfolding S_def by auto qed then have False using \\ P s\ by auto then show ?thesis by simp next case (insertI_2 s0) have *: "\t\{s0<.. P t" using \ \t\{s<.. P t\ that \\t\{s0<.. P t\ by force show ?thesis apply (rule finite_Psegments.insertI_2[of s0]) subgoal using insertI_2.prems(2) local.insertI_2(1) by auto subgoal using \s0 = a \ P s0\ . subgoal using * . subgoal using \finite_Psegments P a s0\ . done qed moreover note \s = b \ P s\ ultimately show ?case by auto qed subsection \Finite segment intersection of a path with the imaginary axis\ definition finite_ReZ_segments::"(real \ complex) \ complex \ bool" where "finite_ReZ_segments g z = finite_Psegments (\t. Re (g t - z) = 0) 0 1" lemma finite_ReZ_segments_joinpaths: assumes g1:"finite_ReZ_segments g1 z" and g2: "finite_ReZ_segments g2 z" and "path g1" "path g2" "pathfinish g1=pathstart g2" shows "finite_ReZ_segments (g1+++g2) z" proof - define P where "P = (\t. (Re ((g1 +++ g2) t - z) = 0 \ 0 t<1) \ t=0 \ t=1)" have "finite_Psegments P 0 (1/2)" proof - have "finite_Psegments (\t. Re (g1 t - z) = 0) 0 1" using g1 unfolding finite_ReZ_segments_def . then have "finite_Psegments (\t. Re (g1 (2 * t) - z) = 0) 0 (1/2)" apply (drule_tac finite_Psegments_pos_linear[of _ 2 0 0 "1/2",simplified]) by (auto simp add:comp_def) then show ?thesis unfolding P_def joinpaths_def by (elim finite_Psegments_congE,auto) qed moreover have "finite_Psegments P (1/2) 1" proof - have "finite_Psegments (\t. Re (g2 t - z) = 0) 0 1" using g2 unfolding finite_ReZ_segments_def . then have "finite_Psegments (\t. Re (g2 (2 * t-1) - z) = 0) (1/2) 1" apply (drule_tac finite_Psegments_pos_linear[of _ 2 "1/2" "-1" 1,simplified]) by (auto simp add:comp_def) then show ?thesis unfolding P_def joinpaths_def apply (elim finite_Psegments_congE) by auto qed moreover have "closed {x. P x}" proof - define Q where "Q=(\t. Re ((g1 +++ g2) t - z) = 0)" have "continuous_on {0<..<1} (g1+++g2)" using path_join_imp[OF \path g1\ \path g2\ \pathfinish g1=pathstart g2\] unfolding path_def by (auto elim:continuous_on_subset) from continuous_on_Re[OF this] have "continuous_on {0<..<1} (\x. Re ((g1 +++ g2) x))" . from continuous_on_open_Collect_neq[OF this,of "\_. Re z",OF continuous_on_const,simplified] have "open {t. Re ((g1 +++ g2) t - z) \ 0 \ 0 t<1}" by (elim arg_elim[where f="open"],auto) from closed_Diff[of "{0::real..1}",OF _ this,simplified] show "closed {x. P x}" apply (elim arg_elim[where f=closed]) by (auto simp add:P_def) qed ultimately have "finite_Psegments P 0 1" using finite_Psegments_combine[of _ 0 "1/2" 1] by auto then show ?thesis unfolding finite_ReZ_segments_def P_def by (elim finite_Psegments_congE,auto) qed lemma finite_ReZ_segments_congE: assumes "finite_ReZ_segments p1 z1" "\t. \0 \ Re(p1 t- z1) = Re(p2 t - z2)" shows "finite_ReZ_segments p2 z2" using assms unfolding finite_ReZ_segments_def apply (elim finite_Psegments_congE) by auto lemma finite_ReZ_segments_constI: assumes "\t. 0t<1 \ g t = c" shows "finite_ReZ_segments g z" proof - have "finite_ReZ_segments (\_. c) z" unfolding finite_ReZ_segments_def by (rule finite_Psegments_constI,auto) then show ?thesis using assms by (elim finite_ReZ_segments_congE,auto) qed lemma finite_ReZ_segment_cases [consumes 1, case_names subEq subNEq,cases pred:finite_ReZ_segments]: assumes "finite_ReZ_segments g z" and subEq:"(\s. \s \ {0..<1};s=0\Re (g s) = Re z; \t\{s<..<1}. Re (g t) = Re z;finite_ReZ_segments (subpath 0 s g) z\ \ P)" and subNEq:"(\s. \s \ {0..<1};s=0\Re (g s) = Re z; \t\{s<..<1}. Re (g t) \ Re z;finite_ReZ_segments (subpath 0 s g) z\ \ P)" shows "P" using assms(1) unfolding finite_ReZ_segments_def proof (cases rule:finite_Psegments.cases) case emptyI then show ?thesis by auto next case (insertI_1 s) have "finite_ReZ_segments (subpath 0 s g) z" proof (cases "s=0") case True show ?thesis apply (rule finite_ReZ_segments_constI) using True unfolding subpath_def by auto next case False then have "s>0" using \s\{0..<1}\ by auto from finite_Psegments_pos_linear[OF _ this,of _ 0 0 1] insertI_1(4) show "finite_ReZ_segments (subpath 0 s g) z" unfolding finite_ReZ_segments_def comp_def subpath_def by auto qed then show ?thesis using subEq insertI_1 by force next case (insertI_2 s) have "finite_ReZ_segments (subpath 0 s g) z" proof (cases "s=0") case True show ?thesis apply (rule finite_ReZ_segments_constI) using True unfolding subpath_def by auto next case False then have "s>0" using \s\{0..<1}\ by auto from finite_Psegments_pos_linear[OF _ this,of _ 0 0 1] insertI_2(4) show "finite_ReZ_segments (subpath 0 s g) z" unfolding finite_ReZ_segments_def comp_def subpath_def by auto qed then show ?thesis using subNEq insertI_2 by force qed lemma finite_ReZ_segments_induct [case_names sub0 subEq subNEq, induct pred:finite_ReZ_segments]: assumes "finite_ReZ_segments g z" assumes sub0:"\g z. (P (subpath 0 0 g) z)" and subEq:"(\s g z. \s \ {0..<1};s=0\Re (g s) = Re z; \t\{s<..<1}. Re (g t) = Re z;finite_ReZ_segments (subpath 0 s g) z; P (subpath 0 s g) z\ \ P g z)" and subNEq:"(\s g z. \s \ {0..<1};s=0\Re (g s) = Re z; \t\{s<..<1}. Re (g t) \ Re z;finite_ReZ_segments (subpath 0 s g) z; P (subpath 0 s g) z\ \ P g z)" shows "P g z" proof - have "finite_Psegments (\t. Re (g t - z) = 0) 0 1" using assms(1) unfolding finite_ReZ_segments_def by auto then have "(0::real)\1 \ P (subpath 0 1 g) z" proof (induct rule: finite_Psegments.induct[of _ 0 1 "\a b. b\a \ P (subpath a b g) z"] ) case (emptyI a b) then show ?case using sub0[of "subpath a b g"] unfolding subpath_def by auto next case (insertI_1 s a b) have ?case when "a=b" using sub0[of "subpath a b g"] that unfolding subpath_def by auto moreover have ?case when "a\b" proof - have "b>a" using that \s \ {a.. by auto define s'::real where "s'=(s-a)/(b-a)" have "P (subpath a b g) z" proof (rule subEq[of s' "subpath a b g"]) show "\t\{s'<..<1}. Re (subpath a b g t) = Re z" proof fix t assume "t \ {s'<..<1}" then have "(b - a) * t + a\{s<..b>a\ \s \ {a.. apply (auto simp add:field_simps) by (sos "((((A<0 * (A<1 * A<2)) * R<1) + (((A<=1 * (A<0 * R<1)) * (R<1 * [1]^2)) + ((A<=0 * (A<0 * (A<1 * R<1))) * (R<1 * [1]^2)))))") then have "Re (g ((b - a) * t + a) - z) = 0" using insertI_1(3)[rule_format,of "(b - a) * t + a"] by auto then show "Re (subpath a b g t) = Re z" unfolding subpath_def by auto qed show "finite_ReZ_segments (subpath 0 s' (subpath a b g)) z" proof (cases "s=a") case True then show ?thesis unfolding s'_def subpath_def by (auto intro:finite_ReZ_segments_constI) next case False have "finite_Psegments (\t. Re (g t - z) = 0) a s" using insertI_1(4) unfolding finite_ReZ_segments_def by auto then have "finite_Psegments ((\t. Re (g t - z) = 0) \ (\t. (s - a) * t + a)) 0 1" apply (elim finite_Psegments_pos_linear[of _ "s-a" 0 a 1,simplified]) using False \s\{a.. by auto then show ?thesis using \b>a\ unfolding finite_ReZ_segments_def subpath_def s'_def comp_def by auto qed show "s' \ {0..<1}" using \b>a\ \s\{a.. unfolding s'_def by (auto simp add:field_simps) show "P (subpath 0 s' (subpath a b g)) z" proof - have "P (subpath a s g) z" using insertI_1(1,5) by auto then show ?thesis using \b>a\ unfolding s'_def subpath_def by simp qed show "s' = 0 \ Re (subpath a b g s') = Re z" proof - have ?thesis when "s=a" using that unfolding s'_def by auto moreover have ?thesis when "Re (g s - z) = 0" using that unfolding s'_def subpath_def by auto ultimately show ?thesis using \s = a \ Re (g s - z) = 0\ by auto qed qed then show ?thesis using \b>a\ by auto qed ultimately show ?case by auto next case (insertI_2 s a b) have ?case when "a=b" using sub0[of "subpath a b g"] that unfolding subpath_def by auto moreover have ?case when "a\b" proof - have "b>a" using that \s \ {a.. by auto define s'::real where "s'=(s-a)/(b-a)" have "P (subpath a b g) z" proof (rule subNEq[of s' "subpath a b g"]) show "\t\{s'<..<1}. Re (subpath a b g t) \ Re z" proof fix t assume "t \ {s'<..<1}" then have "(b - a) * t + a\{s<..b>a\ \s \ {a.. apply (auto simp add:field_simps) by (sos "((((A<0 * (A<1 * A<2)) * R<1) + (((A<=1 * (A<0 * R<1)) * (R<1 * [1]^2)) + ((A<=0 * (A<0 * (A<1 * R<1))) * (R<1 * [1]^2)))))") then have "Re (g ((b - a) * t + a) - z) \ 0" using insertI_2(3)[rule_format,of "(b - a) * t + a"] by auto then show "Re (subpath a b g t) \ Re z" unfolding subpath_def by auto qed show "finite_ReZ_segments (subpath 0 s' (subpath a b g)) z" proof (cases "s=a") case True then show ?thesis unfolding s'_def subpath_def by (auto intro:finite_ReZ_segments_constI) next case False have "finite_Psegments (\t. Re (g t - z) = 0) a s" using insertI_2(4) unfolding finite_ReZ_segments_def by auto then have "finite_Psegments ((\t. Re (g t - z) = 0) \ (\t. (s - a) * t + a)) 0 1" apply (elim finite_Psegments_pos_linear[of _ "s-a" 0 a 1,simplified]) using False \s\{a.. by auto then show ?thesis using \b>a\ unfolding finite_ReZ_segments_def subpath_def s'_def comp_def by auto qed show "s' \ {0..<1}" using \b>a\ \s\{a.. unfolding s'_def by (auto simp add:field_simps) show "P (subpath 0 s' (subpath a b g)) z" proof - have "P (subpath a s g) z" using insertI_2(1,5) by auto then show ?thesis using \b>a\ unfolding s'_def subpath_def by simp qed show "s' = 0 \ Re (subpath a b g s') = Re z" proof - have ?thesis when "s=a" using that unfolding s'_def by auto moreover have ?thesis when "Re (g s - z) = 0" using that unfolding s'_def subpath_def by auto ultimately show ?thesis using \s = a \ Re (g s - z) = 0\ by auto qed qed then show ?thesis using \b>a\ by auto qed ultimately show ?case by auto qed then show ?thesis by auto qed lemma finite_ReZ_segments_shiftpah: assumes "finite_ReZ_segments g z" "s\{0..1}" "path g" and loop:"pathfinish g = pathstart g" shows "finite_ReZ_segments (shiftpath s g) z" proof - have "finite_Psegments (\t. Re (shiftpath s g t - z) = 0) 0 (1-s)" proof - have "finite_Psegments (\t. Re (g t) = Re z) s 1" using assms finite_Psegments_included[of _ 0 1 s] unfolding finite_ReZ_segments_def by force then have "finite_Psegments (\t. Re (g (s + t) - z) = 0) 0 (1-s)" using finite_Psegments_pos_linear[of "\t. Re (g t - z) =0" 1 0 s "1-s",simplified] unfolding comp_def by (auto simp add:algebra_simps) then show ?thesis unfolding shiftpath_def apply (elim finite_Psegments_congE) using \s\{0..1}\ by auto qed moreover have "finite_Psegments (\t. Re (shiftpath s g t - z) = 0) (1-s) 1" proof - have "finite_Psegments (\t. Re (g t) = Re z) 0 s " using assms finite_Psegments_included unfolding finite_ReZ_segments_def by force then have "finite_Psegments (\t. Re (g (s + t - 1) - z) = 0) (1-s) 1" using finite_Psegments_pos_linear[of "\t. Re (g t - z) =0" 1 "1-s" "s-1" 1,simplified] unfolding comp_def by (auto simp add:algebra_simps) then show ?thesis unfolding shiftpath_def apply (elim finite_Psegments_congE) using \s\{0..1}\ by auto qed moreover have "1 - s \ {0..1}" using \s\{0..1}\ by auto moreover have "closed ({x. Re (shiftpath s g x - z) = 0} \ {0..1})" proof - let ?f = "\x. Re (shiftpath s g x - z)" have "continuous_on {0..1} ?f" using path_shiftpath[OF \path g\ loop \s\{0..1}\] unfolding path_def by (auto intro: continuous_intros) from continuous_closed_preimage_constant[OF this,of 0,simplified] show ?thesis apply (elim arg_elim[of closed]) by force qed ultimately show ?thesis unfolding finite_ReZ_segments_def by (rule finite_Psegments_combine[where b="1-s"]) qed lemma finite_imp_finite_ReZ_segments: assumes "finite {t. Re (g t - z) = 0 \ 0 \ t \ t\1}" shows "finite_ReZ_segments g z" proof - define P where "P = (\t. Re (g t - z) = 0)" define rs where "rs=(\b. {t. P t \ 0 < t \ t0" for b using that proof (induct "card (rs b)" arbitrary:b rule:nat_less_induct) case ind:1 have ?case when "rs b= {}" apply (rule finite_Psegments.intros(3)[of 0]) using that \0 < b\ unfolding rs_def by (auto intro:finite_Psegments.intros) moreover have ?case when "rs b\{}" proof - define lj where "lj = Max (rs b)" have "0finite (rs b)\ \rs b\{}\,folded lj_def] unfolding rs_def by auto show ?thesis proof (rule finite_Psegments.intros(3)[of lj]) show "lj \ {0.. P lj" using \0 \lj \P lj\ by auto show "\t\{lj<.. P t" proof (rule ccontr) assume " \ (\t\{lj<.. P t) " then obtain t where t:"P t" "lj < t" "t < b" by auto then have "t\rs b" unfolding rs_def using \lj>0\ by auto then have "t\lj" using Max_ge[OF \finite (rs b)\,of t] unfolding lj_def by auto then show False using \t>lj\ by auto qed show "finite_Psegments P 0 lj" proof (rule ind.hyps[rule_format,of "card (rs lj)" lj,simplified]) show "finite (rs lj)" using \finite (rs b)\ unfolding rs_def using \lj by (auto elim!:rev_finite_subset ) show "card (rs lj) < card (rs b)" apply (rule psubset_card_mono[OF \finite (rs b)\]) using Max_in \finite (rs lj)\ \lj < b\ lj_def rs_def that by fastforce show "0 < lj" using \0 . qed qed qed ultimately show ?case by auto qed moreover have "finite (rs 1)" using assms unfolding rs_def P_def by (auto elim:rev_finite_subset) ultimately have "finite_Psegments P 0 1" by auto then show ?thesis unfolding P_def finite_ReZ_segments_def . qed lemma finite_ReZ_segments_poly_linepath: shows "finite_ReZ_segments (poly p o linepath a b) z" proof - define P where "P=map_poly Re (pcompose (p-[:z:]) [:a,b-a:])" have *:"Re ((poly p \ linepath a b) t - z) = 0 \ poly P t=0" for t unfolding inner_complex_def P_def linepath_def comp_def apply (subst Re_poly_of_real[symmetric]) by (auto simp add: algebra_simps poly_pcompose scaleR_conv_of_real) have ?thesis when "P\0" proof - have "finite {t. poly P t=0}" using that poly_roots_finite by auto then have "finite {t. Re ((poly p \ linepath a b) t - z) = 0 \ 0 \ t \ t \ 1}" using * by auto then show ?thesis using finite_imp_finite_ReZ_segments[of "poly p o linepath a b" z] by auto qed moreover have ?thesis when "P=0" unfolding finite_ReZ_segments_def apply (rule finite_Psegments_constI[where c=True]) apply (subst *) using that by auto ultimately show ?thesis by auto qed lemma part_circlepath_half_finite_inter: assumes "st\tt" "r\0" "c\0" shows "finite {t. part_circlepath z0 r st tt t \ c = d \ 0\ t \ t\1}" (is "finite ?T") proof - let ?S = "{\. (z0+r*exp (\ * \ )) \ c = d \ \ \ closed_segment st tt}" define S where "S \ {\. (z0+r*exp (\ * \ )) \ c = d \ \ \ closed_segment st tt}" have "S = linepath st tt ` ?T" proof define g where "g\(\t. (t-st)/(tt -st))" have "0\g t" "g t\1" when "t \ closed_segment st tt " for t using that \st\tt\ closed_segment_eq_real_ivl unfolding g_def real_scaleR_def by (auto simp add:divide_simps) moreover have "linepath st tt (g t) =t" "g (linepath st tt t) = t" for t unfolding linepath_def g_def real_scaleR_def using \st\tt\ apply (simp_all add:divide_simps) by (auto simp add:algebra_simps ) ultimately have "x\linepath st tt ` ?T" when "x\S" for x using that unfolding S_def by (auto intro!:image_eqI[where x="g x"] simp add:part_circlepath_def) then show "S \ linepath st tt ` ?T" by auto next have "x\S" when "x\linepath st tt ` ?T" for x using that unfolding part_circlepath_def S_def by (auto simp add: linepath_in_path) then show "linepath st tt ` ?T \ S" by auto qed moreover have "finite S" proof - define a' b' c' where "a'=r * Re c" and "b' = r* Im c" and "c'=Im c * Im z0 + Re z0 * Re c - d" define f where "f \= a' * cos \ + b' * sin \ + c'" for \ have "(z0+r*exp (\ * \ )) \ c = d \ f \ = 0" for \ unfolding exp_Euler inner_complex_def f_def a'_def b'_def c'_def by (auto simp add:algebra_simps cos_of_real sin_of_real) then have *:"S = roots f \ closed_segment st tt" unfolding S_def roots_within_def by auto have "uniform_discrete S" proof - have "a' \ 0 \ b' \ 0 \ c' \ 0" using assms complex_eq_iff unfolding a'_def b'_def c'_def by auto then have "periodic_set (roots f) (4 * pi)" using periodic_set_sin_cos_linear[of a' b' c',folded f_def] by auto then have "uniform_discrete (roots f)" using periodic_imp_uniform_discrete by auto then show ?thesis unfolding * by auto qed moreover have "bounded S" unfolding * by (simp add: bounded_Int bounded_closed_segment) ultimately show ?thesis using uniform_discrete_finite_iff by auto qed moreover have "inj_on (linepath st tt) ?T" proof - have "inj (linepath st tt)" unfolding linepath_def using assms inj_segment by blast then show ?thesis by (auto elim:subset_inj_on) qed ultimately show ?thesis by (auto elim!: finite_imageD) qed lemma linepath_half_finite_inter: assumes "a \ c \ d \ b \ c \ d" shows "finite {t. linepath a b t \ c = d \ 0\ t \ t\1}" (is "finite ?S") proof (rule ccontr) assume asm:"infinite ?S" obtain t1 t2 where u1u2:"t1\t2" "t1\?S" "t2\?S" proof - obtain t1 where "t1\?S" using not_finite_existsD asm by blast moreover have "\u2. u2\?S-{t1}" using infinite_remove[OF asm,of t1] by (meson finite.emptyI rev_finite_subset subsetI) ultimately show ?thesis using that by auto qed have t1:"(1-t1)*(a \ c) + t1 * (b \ c) = d" using \t1\?S\ unfolding linepath_def by (simp add: inner_left_distrib) have t2:"(1-t2)*(a \ c) + t2 * (b \ c) = d" using \t2\?S\ unfolding linepath_def by (simp add: inner_left_distrib) have "a \ c = d" proof - have "t2*((1-t1)*(a \ c) + t1 * (b \ c)) = t2*d" using t1 by auto then have *:"(t2-t1*t2)*(a \ c) + t1*t2 * (b \ c) = t2*d" by (auto simp add:algebra_simps) have "t1*((1-t2)*(a \ c) + t2 * (b \ c)) = t1*d" using t2 by auto then have **:"(t1-t1*t2)*(a \ c) + t1*t2 * (b \ c) = t1*d" by (auto simp add:algebra_simps) have "(t2-t1)*(a \ c) = (t2-t1)*d" using arg_cong2[OF * **,of minus] by (auto simp add:algebra_simps) then show ?thesis using \t1\t2\ by auto qed moreover have "b \ c = d" proof - have "(1-t2)*((1-t1)*(a \ c) + t1 * (b \ c)) = (1-t2)*d" using t1 by auto then have *:"(1-t1)*(1-t2)*(a \ c) + (t1-t1*t2) * (b \ c) = (1-t2)*d" by (auto simp add:algebra_simps) have "(1-t1)*((1-t2)*(a \ c) + t2 * (b \ c)) = (1-t1)*d" using t2 by auto then have **:"(1-t1)*(1-t2)*(a \ c) + (t2-t1*t2) * (b \ c) = (1-t1)*d" by (auto simp add:algebra_simps) have "(t2-t1)*(b \ c) = (t2-t1)*d" using arg_cong2[OF ** *,of minus] by (auto simp add:algebra_simps) then show ?thesis using \t1\t2\ by auto qed ultimately show False using assms by auto qed lemma finite_half_joinpaths_inter: assumes "finite {t. l1 t \ c = d \ 0\ t \ t\1}" "finite {t. l2 t \ c = d \ 0\ t \ t\1}" shows "finite {t. (l1+++l2) t \ c = d \ 0\ t \ t\1}" proof - let ?l1s = "{t. l1 (2*t) \ c = d \ 0\ t \ t\1/2}" let ?l2s = "{t. l2 (2 * t - 1) \ c = d \ 1/2< t \ t\1}" let ?ls = "\l. {t. l t \ c = d \ 0\ t \ t\1}" have "{t. (l1+++l2) t \ c = d \ 0\ t \ t\1} = ?l1s \ ?l2s" unfolding joinpaths_def by auto moreover have "finite ?l1s" proof - have "?l1s = ((*) (1/2)) ` ?ls l1" by (auto intro:rev_image_eqI) thus ?thesis using assms by simp qed moreover have "finite ?l2s" proof - have "?l2s \ (\x. x/2 + 1/2) ` ?ls l2" by (auto intro:rev_image_eqI simp add:field_simps) thus ?thesis using assms by (auto elim:finite_subset) qed ultimately show ?thesis by simp qed lemma finite_ReZ_segments_linepath: "finite_ReZ_segments (linepath a b) z" proof - have ?thesis when "Re a\Re z \ Re b \Re z" proof - let ?S1="{t. Re (linepath a b t-z) = 0 \ 0 \ t \ t \ 1}" have "finite ?S1" using linepath_half_finite_inter[of a "Complex 1 0" "Re z" b] that using one_complex.code by auto from finite_imp_finite_ReZ_segments[OF this] show ?thesis . qed moreover have ?thesis when "Re a=Re z" "Re b=Re z" unfolding finite_ReZ_segments_def apply (rule finite_Psegments.intros(2)[of 0]) using that unfolding linepath_def by (auto simp add:algebra_simps intro:finite_Psegments.intros) ultimately show ?thesis by blast qed lemma finite_ReZ_segments_part_circlepath: "finite_ReZ_segments (part_circlepath z0 r st tt) z" proof - have ?thesis when "st \ tt" "r \ 0" proof - let ?S1="{t. Re (part_circlepath z0 r st tt t-z) = 0 \ 0 \ t \ t \ 1}" have "finite ?S1" using part_circlepath_half_finite_inter[of st tt r "Complex 1 0" z0 "Re z"] that one_complex.code by (auto simp add:inner_complex_def ) from finite_imp_finite_ReZ_segments[OF this] show ?thesis . qed moreover have ?thesis when "st =tt \ r=0" proof - define c where "c = z0 + r * exp (\ * tt)" have "part_circlepath z0 r st tt = (\t. c)" unfolding part_circlepath_def c_def using that linepath_refl by auto then show ?thesis using finite_ReZ_segments_linepath[of c c z] linepath_refl[of c] by auto qed ultimately show ?thesis by blast qed lemma finite_ReZ_segments_poly_of_real: shows "finite_ReZ_segments (poly p o of_real) z" using finite_ReZ_segments_poly_linepath[of p 0 1 z] unfolding linepath_def by (auto simp add:scaleR_conv_of_real) lemma finite_ReZ_segments_subpath: assumes "finite_ReZ_segments g z" "0\u" "u\v" "v\1" shows "finite_ReZ_segments (subpath u v g) z" proof (cases "u=v") case True then show ?thesis unfolding subpath_def by (auto intro:finite_ReZ_segments_constI) next case False then have "uu\v\ by auto define P where "P=(\t. Re (g t - z) = 0)" have "finite_ReZ_segments (subpath u v g) z = finite_Psegments (P o (\t. (v - u) * t + u)) 0 1" unfolding finite_ReZ_segments_def subpath_def P_def comp_def by auto also have "..." apply (rule finite_Psegments_pos_linear) using assms False unfolding finite_ReZ_segments_def by (fold P_def,auto elim:finite_Psegments_included) finally show ?thesis . qed subsection \jump and jumpF\ definition jump::"(real \ real) \ real \ int" where "jump f a = (if (LIM x (at_left a). f x :> at_bot) \ (LIM x (at_right a). f x :> at_top) then 1 else if (LIM x (at_left a). f x :> at_top) \ (LIM x (at_right a). f x :> at_bot) then -1 else 0)" definition jumpF::"(real \ real) \ real filter \ real" where "jumpF f F \ (if filterlim f at_top F then 1/2 else if filterlim f at_bot F then -1/2 else (0::real))" lemma jumpF_const[simp]: assumes "F\bot" shows "jumpF (\_. c) F = 0" proof - have False when "LIM x F. c :> at_bot" using filterlim_at_bot_nhds[OF that _ \F\bot\] by auto moreover have False when "LIM x F. c :> at_top" using filterlim_at_top_nhds[OF that _ \F\bot\] by auto ultimately show ?thesis unfolding jumpF_def by auto qed lemma jumpF_not_infinity: assumes "continuous F g" "F\bot" shows "jumpF g F = 0" proof - have "\ filterlim g at_infinity F" using not_tendsto_and_filterlim_at_infinity[OF \F \bot\ assms(1)[unfolded continuous_def]] by auto then have "\ filterlim g at_bot F" "\ filterlim g at_top F" using at_bot_le_at_infinity at_top_le_at_infinity filterlim_mono by blast+ then show ?thesis unfolding jumpF_def by auto qed lemma jumpF_linear_comp: assumes "c\0" shows "jumpF (f o (\x. c*x+b)) (at_left x) = (if c>0 then jumpF f (at_left (c*x+b)) else jumpF f (at_right (c*x+b)))" (is ?case1) "jumpF (f o (\x. c*x+b)) (at_right x) = (if c>0 then jumpF f (at_right (c*x+b)) else jumpF f (at_left (c*x+b)))" (is ?case2) proof - let ?g = "\x. c*x+b" have ?case1 ?case2 when "\ c>0" proof - have "c<0" using \c\0\ that by auto have "filtermap ?g (at_left x) = at_right (?g x)" "filtermap ?g (at_right x) = at_left (?g x)" using \c<0\ filtermap_linear_at_left[OF \c\0\, of b x] filtermap_linear_at_right[OF \c\0\, of b x] by auto then have "jumpF (f \ ?g) (at_left x) = jumpF f (at_right (?g x))" "jumpF (f \ ?g) (at_right x) = jumpF f (at_left (?g x))" unfolding jumpF_def filterlim_def comp_def by (auto simp add: filtermap_filtermap[of f ?g,symmetric]) then show ?case1 ?case2 using \c<0\ by auto qed moreover have ?case1 ?case2 when "c>0" proof - have "filtermap ?g (at_left x) = at_left (?g x)" "filtermap ?g (at_right x) = at_right (?g x)" using that filtermap_linear_at_left[OF \c\0\, of b x] filtermap_linear_at_right[OF \c\0\, of b x] by auto then have "jumpF (f \ ?g) (at_left x) = jumpF f (at_left (?g x))" "jumpF (f \ ?g) (at_right x) = jumpF f (at_right (?g x))" unfolding jumpF_def filterlim_def comp_def by (auto simp add: filtermap_filtermap[of f ?g,symmetric]) then show ?case1 ?case2 using that by auto qed ultimately show ?case1 ?case2 by auto qed lemma jump_const[simp]:"jump (\_. c) a = 0" proof - have False when "LIM x (at_left a). c :> at_bot" apply (rule not_tendsto_and_filterlim_at_infinity[of "at_left a" "\_. c" c]) apply auto using at_bot_le_at_infinity filterlim_mono that by blast moreover have False when "LIM x (at_left a). c :> at_top" apply (rule not_tendsto_and_filterlim_at_infinity[of "at_left a" "\_. c" c]) apply auto using at_top_le_at_infinity filterlim_mono that by blast ultimately show ?thesis unfolding jump_def by auto qed lemma jump_not_infinity: "isCont f a \ jump f a =0" by (meson at_bot_le_at_infinity at_top_le_at_infinity filterlim_at_split filterlim_def isCont_def jump_def not_tendsto_and_filterlim_at_infinity order_trans trivial_limit_at_left_real) lemma jump_jump_poly_aux: assumes "p\0" "coprime p q" shows "jump (\x. poly q x / poly p x) a = jump_poly q p a" proof (cases "q=0") case True then show ?thesis by auto next case False define f where "f \ (\x. poly q x / poly p x)" have ?thesis when "poly q a = 0" proof - have "poly p a\0" using coprime_poly_0[OF \coprime p q\] that by blast then have "isCont f a" unfolding f_def by simp then have "jump f a=0" using jump_not_infinity by auto moreover have "jump_poly q p a=0" using jump_poly_not_root[OF \poly p a\0\] by auto ultimately show ?thesis unfolding f_def by auto qed moreover have ?thesis when "poly q a\0" proof (cases "even(order a p)") case True define c where "c\sgn (poly q a)" note filterlim_divide_at_bot_at_top_iff [OF _ that,of "poly q" "at_left a" "poly p",folded f_def c_def,simplified] filterlim_divide_at_bot_at_top_iff [OF _ that,of "poly q" "at_right a" "poly p",folded f_def c_def,simplified] moreover have "(poly p has_sgnx - c) (at_left a) = (poly p has_sgnx - c) (at_right a)" "(poly p has_sgnx c) (at_left a) = (poly p has_sgnx c) (at_right a)" using poly_has_sgnx_left_right[OF \p\0\] True by auto moreover have "c\0" by (simp add: c_def sgn_if that) then have False when "(poly p has_sgnx - c) (at_right a)" "(poly p has_sgnx c) (at_right a)" using has_sgnx_unique[OF _ that] by auto ultimately have "jump f a = 0" unfolding jump_def by auto moreover have "jump_poly q p a = 0" unfolding jump_poly_def using True by (simp add: order_0I that) ultimately show ?thesis unfolding f_def by auto next case False define c where "c\sgn (poly q a)" have "(poly p \ 0) (at a)" using False by (metis even_zero order_0I poly_tendsto(1)) then have "(poly p \ 0) (at_left a)" and "(poly p \ 0) (at_right a)" by (auto simp add: filterlim_at_split) moreover note filterlim_divide_at_bot_at_top_iff [OF _ that,of "poly q" _ "poly p",folded f_def c_def] moreover have "(poly p has_sgnx c) (at_left a) = (poly p has_sgnx - c) (at_right a)" "(poly p has_sgnx - c) (at_left a) = (poly p has_sgnx c) (at_right a)" using poly_has_sgnx_left_right[OF \p\0\] False by auto ultimately have "jump f a = (if (poly p has_sgnx c) (at_right a) then 1 else if (poly p has_sgnx - c) (at_right a) then -1 else 0)" unfolding jump_def by auto also have "... = (if sign_r_pos (q * p) a then 1 else - 1)" proof - have "(poly p has_sgnx c) (at_right a) \ sign_r_pos (q * p) a " proof assume "(poly p has_sgnx c) (at_right a)" then have "sgnx (poly p) (at_right a) = c" by auto moreover have "sgnx (poly q) (at_right a) = c" unfolding c_def using that by (auto intro!: tendsto_nonzero_sgnx) ultimately have "sgnx (\x. poly (q*p) x) (at_right a) = c * c" by (simp add:sgnx_times) moreover have "c\0" by (simp add: c_def sgn_if that) ultimately have "sgnx (\x. poly (q*p) x) (at_right a) > 0" using not_real_square_gt_zero by fastforce then show "sign_r_pos (q * p) a" using sign_r_pos_sgnx_iff by blast next assume asm:"sign_r_pos (q * p) a" let ?c1 = "sgnx (poly p) (at_right a)" let ?c2 = "sgnx (poly q) (at_right a)" have "0 < sgnx (\x. poly (q * p) x) (at_right a)" using asm sign_r_pos_sgnx_iff by blast then have "?c2 * ?c1 >0" apply (subst (asm) poly_mult) apply (subst (asm) sgnx_times) by auto then have "?c2>0 \ ?c1>0 \ ?c2<0 \ ?c1<0" by (simp add: zero_less_mult_iff) then have "?c1=?c2" using sgnx_values[OF sgnx_able_poly(1), of a,simplified] by (metis add.inverse_neutral less_minus_iff less_not_sym) moreover have "sgnx (poly q) (at_right a) = c" unfolding c_def using that by (auto intro!: tendsto_nonzero_sgnx) ultimately have "?c1 = c" by auto then show "(poly p has_sgnx c) (at_right a)" using sgnx_able_poly(1) sgnx_able_sgnx by blast qed then show ?thesis unfolding jump_poly_def using poly_has_sgnx_values[OF \p\0\] by (metis add.inverse_inverse c_def sgn_if that) qed also have "... = jump_poly q p a" unfolding jump_poly_def using False order_root that by (simp add: order_root assms(1)) finally show ?thesis unfolding f_def by auto qed ultimately show ?thesis by auto qed lemma jump_jumpF: assumes cont:"isCont (inverse o f) a" and sgnxl:"(f has_sgnx l) (at_left a)" and sgnxr:"(f has_sgnx r) (at_right a)" and "l\0 " "r\0" shows "jump f a = jumpF f (at_right a) - jumpF f (at_left a)" proof - have ?thesis when "filterlim f at_bot (at_left a)" "filterlim f at_top (at_right a)" unfolding jump_def jumpF_def using that filterlim_at_top_at_bot[OF _ _ trivial_limit_at_left_real] by auto moreover have ?thesis when "filterlim f at_top (at_left a)" "filterlim f at_bot (at_right a)" unfolding jump_def jumpF_def using that filterlim_at_top_at_bot[OF _ _ trivial_limit_at_right_real] by auto moreover have ?thesis when "\ filterlim f at_bot (at_left a) \ \ filterlim f at_top (at_right a)" "\ filterlim f at_top (at_left a) \ \ filterlim f at_bot (at_right a)" proof (cases "f a=0") case False have "jumpF f (at_right a) = 0" "jumpF f (at_left a) = 0" proof - have "isCont (inverse o inverse o f) a" using cont False unfolding comp_def by (rule_tac continuous_at_within_inverse, auto) then have "isCont f a" unfolding comp_def by auto then have "(f \ f a) (at_right a)" "(f \ f a) (at_left a)" unfolding continuous_at_split by (auto simp add:continuous_within) moreover note trivial_limit_at_left_real trivial_limit_at_right_real ultimately show "jumpF f (at_right a) = 0" "jumpF f (at_left a) = 0" unfolding jumpF_def using filterlim_at_bot_nhds filterlim_at_top_nhds by metis+ qed then show ?thesis unfolding jump_def using that by auto next case True then have tends0:"((\x. inverse (f x)) \ 0) (at a)" using cont unfolding isCont_def comp_def by auto have "jump f a = 0" using that unfolding jump_def by auto have r_lim:"if r>0 then filterlim f at_top (at_right a) else filterlim f at_bot (at_right a)" proof (cases "r>0") case True then have "\\<^sub>F x in (at_right a). 0 < f x" using sgnxr unfolding has_sgnx_def by (auto elim:eventually_mono) then have "filterlim f at_top (at_right a)" using filterlim_inverse_at_top[of "\x. inverse (f x)", simplified] tends0 unfolding filterlim_at_split by auto then show ?thesis using True by presburger next case False then have "\\<^sub>F x in (at_right a). 0 > f x" using sgnxr \r\0\ False unfolding has_sgnx_def apply (elim eventually_mono) by (meson linorder_neqE_linordered_idom sgn_less) then have "filterlim f at_bot (at_right a)" using filterlim_inverse_at_bot[of "\x. inverse (f x)", simplified] tends0 unfolding filterlim_at_split by auto then show ?thesis using False by simp qed have l_lim:"if l>0 then filterlim f at_top (at_left a) else filterlim f at_bot (at_left a)" proof (cases "l>0") case True then have "\\<^sub>F x in (at_left a). 0 < f x" using sgnxl unfolding has_sgnx_def by (auto elim:eventually_mono) then have "filterlim f at_top (at_left a)" using filterlim_inverse_at_top[of "\x. inverse (f x)", simplified] tends0 unfolding filterlim_at_split by auto then show ?thesis using True by presburger next case False then have "\\<^sub>F x in (at_left a). 0 > f x" using sgnxl \l\0\ False unfolding has_sgnx_def apply (elim eventually_mono) by (meson linorder_neqE_linordered_idom sgn_less) then have "filterlim f at_bot (at_left a)" using filterlim_inverse_at_bot[of "\x. inverse (f x)", simplified] tends0 unfolding filterlim_at_split by auto then show ?thesis using False by simp qed have ?thesis when "l>0" "r>0" using that l_lim r_lim \jump f a=0\ unfolding jumpF_def by auto moreover have ?thesis when "\ l>0" "\ r>0" proof - have "filterlim f at_bot (at_right a)" "filterlim f at_bot (at_left a)" using r_lim l_lim that by auto moreover then have "\ filterlim f at_top (at_right a)" "\ filterlim f at_top (at_left a)" by (auto elim: filterlim_at_top_at_bot) ultimately have "jumpF f (at_right a) = -1/2 " "jumpF f (at_left a) = -1/2" unfolding jumpF_def by auto then show ?thesis using \jump f a=0\ by auto qed moreover have ?thesis when "l>0" "\ r>0" proof - note \\ filterlim f at_top (at_left a) \ \ filterlim f at_bot (at_right a)\ moreover have "filterlim f at_bot (at_right a)" "filterlim f at_top (at_left a)" using r_lim l_lim that by auto ultimately have False by auto then show ?thesis by auto qed moreover have ?thesis when "\ l>0" "r>0" proof - note \\ filterlim f at_bot (at_left a) \ \ filterlim f at_top (at_right a)\ moreover have "filterlim f at_bot (at_left a)" "filterlim f at_top (at_right a)" using r_lim l_lim that by auto ultimately have False by auto then show ?thesis by auto qed ultimately show ?thesis by auto qed ultimately show ?thesis by auto qed lemma jump_linear_comp: assumes "c\0" shows "jump (f o (\x. c*x+b)) x = (if c>0 then jump f (c*x+b) else -jump f (c*x+b))" proof (cases "c>0") case False then have "c<0" using \c\0\ by auto let ?g = "\x. c*x+b" have "filtermap ?g (at_left x) = at_right (?g x)" "filtermap ?g (at_right x) = at_left (?g x)" using \c<0\ filtermap_linear_at_left[OF \c\0\, of b x] filtermap_linear_at_right[OF \c\0\, of b x] by auto then have "jump (f \ ?g) x = - jump f (c * x + b)" unfolding jump_def filterlim_def comp_def apply (auto simp add: filtermap_filtermap[of f ?g,symmetric]) apply (fold filterlim_def) by (auto elim:filterlim_at_top_at_bot) then show ?thesis using \c<0\ by auto next case True let ?g = "\x. c*x+b" have "filtermap ?g (at_left x) = at_left (?g x)" "filtermap ?g (at_right x) = at_right (?g x)" using True filtermap_linear_at_left[OF \c\0\, of b x] filtermap_linear_at_right[OF \c\0\, of b x] by auto then have "jump (f \ ?g) x = jump f (c * x + b)" unfolding jump_def filterlim_def comp_def by (auto simp add: filtermap_filtermap[of f ?g,symmetric]) then show ?thesis using True by auto qed lemma jump_divide_derivative: assumes "isCont f x" "g x = 0" "f x\0" and g_deriv:"(g has_field_derivative c) (at x)" and "c\0" shows "jump (\t. f t/g t) x = (if sgn c = sgn ( f x) then 1 else -1)" proof - have g_tendsto:"(g \ 0) (at_left x)" "(g \ 0) (at_right x)" by (metis DERIV_isCont Lim_at_imp_Lim_at_within assms(2) assms(4) continuous_at)+ have f_tendsto:"(f \ f x) (at_left x)" "(f \ f x) (at_right x)" using Lim_at_imp_Lim_at_within assms(1) continuous_at by blast+ have ?thesis when "c>0" "f x>0" proof - have "(g has_sgnx - sgn (f x)) (at_left x)" using has_sgnx_derivative_at_left[OF g_deriv \g x=0\] that by auto moreover have "(g has_sgnx sgn (f x)) (at_right x)" using has_sgnx_derivative_at_right[OF g_deriv \g x=0\] that by auto ultimately have "(LIM t at_left x. f t / g t :> at_bot) \ (LIM t at_right x. f t / g t :> at_top)" using filterlim_divide_at_bot_at_top_iff[OF _ \f x\0\, of f] using f_tendsto(1) f_tendsto(2) g_tendsto(1) g_tendsto(2) by blast moreover have "sgn c = sgn (f x)" using that by auto ultimately show ?thesis unfolding jump_def by auto qed moreover have ?thesis when "c>0" "f x<0" proof - have "(g has_sgnx sgn (f x)) (at_left x)" using has_sgnx_derivative_at_left[OF g_deriv \g x=0\] that by auto moreover have "(g has_sgnx - sgn (f x)) (at_right x)" using has_sgnx_derivative_at_right[OF g_deriv \g x=0\] that by auto ultimately have "(LIM t at_left x. f t / g t :> at_top) \ (LIM t at_right x. f t / g t :> at_bot)" using filterlim_divide_at_bot_at_top_iff[OF _ \f x\0\, of f] using f_tendsto(1) f_tendsto(2) g_tendsto(1) g_tendsto(2) by blast moreover from this have "\ (LIM t at_left x. f t / g t :> at_bot)" using filterlim_at_top_at_bot by fastforce moreover have "sgn c \ sgn (f x)" using that by auto ultimately show ?thesis unfolding jump_def by auto qed moreover have ?thesis when "c<0" "f x>0" proof - have "(g has_sgnx sgn (f x)) (at_left x)" using has_sgnx_derivative_at_left[OF g_deriv \g x=0\] that by auto moreover have "(g has_sgnx - sgn (f x)) (at_right x)" using has_sgnx_derivative_at_right[OF g_deriv \g x=0\] that by auto ultimately have "(LIM t at_left x. f t / g t :> at_top) \ (LIM t at_right x. f t / g t :> at_bot)" using filterlim_divide_at_bot_at_top_iff[OF _ \f x\0\, of f] using f_tendsto(1) f_tendsto(2) g_tendsto(1) g_tendsto(2) by blast moreover from this have "\ (LIM t at_left x. f t / g t :> at_bot)" using filterlim_at_top_at_bot by fastforce moreover have "sgn c \ sgn (f x)" using that by auto ultimately show ?thesis unfolding jump_def by auto qed moreover have ?thesis when "c<0" "f x<0" proof - have "(g has_sgnx - sgn (f x)) (at_left x)" using has_sgnx_derivative_at_left[OF g_deriv \g x=0\] that by auto moreover have "(g has_sgnx sgn (f x)) (at_right x)" using has_sgnx_derivative_at_right[OF g_deriv \g x=0\] that by auto ultimately have "(LIM t at_left x. f t / g t :> at_bot) \ (LIM t at_right x. f t / g t :> at_top)" using filterlim_divide_at_bot_at_top_iff[OF _ \f x\0\, of f] using f_tendsto(1) f_tendsto(2) g_tendsto(1) g_tendsto(2) by blast moreover have "sgn c =sgn (f x)" using that by auto ultimately show ?thesis unfolding jump_def by auto qed ultimately show ?thesis using \c\0\ \f x\0\ by argo qed lemma jump_jump_poly: "jump (\x. poly q x / poly p x) a = jump_poly q p a" proof (cases "p=0") case True then show ?thesis by auto next case False obtain p' q' where p':"p= p'*gcd p q" and q':"q=q'*gcd p q" using gcd_dvd1 gcd_dvd2 dvd_def[of "gcd p q", simplified mult.commute] by metis then have "coprime p' q'" "p'\0" "gcd p q\0" using gcd_coprime \p\0\ by auto define f where "f \ (\x. poly q' x / poly p' x)" define g where "g \ (\x. if poly (gcd p q) x = 0 then 0::real else 1)" have g_tendsto:"(g \ 1) (at_left a)" "(g \ 1) (at_right a)" proof - have "(poly (gcd p q) has_sgnx 1) (at_left a) \ (poly (gcd p q) has_sgnx - 1) (at_left a)" "(poly (gcd p q) has_sgnx 1) (at_right a) \ (poly (gcd p q) has_sgnx - 1) (at_right a)" using \p\0\ poly_has_sgnx_values by auto then have " \\<^sub>F x in at_left a. g x = 1" " \\<^sub>F x in at_right a. g x = 1" unfolding has_sgnx_def g_def by (auto elim:eventually_mono) then show "(g \ 1) (at_left a)" "(g \ 1) (at_right a)" using tendsto_eventually by auto qed have "poly q x / poly p x = g x * f x" for x unfolding f_def g_def by (subst p',subst q',auto) then have "jump (\x. poly q x / poly p x) a = jump (\x. g x * f x) a" by auto also have "... = jump f a" unfolding jump_def apply (subst (1 2) filterlim_tendsto_pos_mult_at_top_iff) prefer 5 apply (subst (1 2) filterlim_tendsto_pos_mult_at_bot_iff) using g_tendsto by auto also have "... = jump_poly q' p' a" using jump_jump_poly_aux[OF \p'\0\ \coprime p' q'\] unfolding f_def by auto also have "... = jump_poly q p a" using jump_poly_mult[OF \gcd p q \ 0\, of q'] p' q' by (metis mult.commute) finally show ?thesis . qed lemma jump_Im_divide_Re_0: assumes "path g" "Re (g x)\0" "0t. Im (g t) / Re (g t)) x = 0" proof - have "isCont g x" using \path g\[unfolded path_def] \0 \x<1\ apply (elim continuous_on_interior) by auto then have "isCont (\t. Im(g t)/Re(g t)) x" using \Re (g x)\0\ by (auto intro:continuous_intros isCont_Re isCont_Im) then show "jump (\t. Im(g t)/Re(g t)) x=0" using jump_not_infinity by auto qed lemma jumpF_im_divide_Re_0: assumes "path g" "Re (g x)\0" shows "\0\x;x<1\ \ jumpF (\t. Im (g t) / Re (g t)) (at_right x) = 0" "\01\ \ jumpF (\t. Im (g t) / Re (g t)) (at_left x) = 0" proof - define g' where "g' = (\t. Im (g t) / Re (g t))" show "jumpF g' (at_right x) = 0" when "0\x" "x<1" proof - have "(g' \ g' x) (at_right x)" proof (cases "x=0") case True have "continuous (at_right 0) g" using \path g\ unfolding path_def by (auto elim:continuous_on_at_right) then have "continuous (at_right x) (\t. Im(g t))" "continuous (at_right x) (\t. Re(g t))" using continuous_Im continuous_Re True by auto moreover have "Re (g (netlimit (at_right x))) \ 0" using assms(2) by (simp add: Lim_ident_at) ultimately have "continuous (at_right x) (\t. Im (g t)/Re(g t))" by (auto intro:continuous_divide) then show ?thesis unfolding g'_def continuous_def by (simp add: Lim_ident_at) next case False have "isCont (\x. Im (g x)) x" "isCont (\x. Re (g x)) x" using \path g\ unfolding path_def by (metis False atLeastAtMost_iff at_within_Icc_at continuous_Im continuous_Re continuous_on_eq_continuous_within less_le that)+ then have "isCont g' x" using assms(2) unfolding g'_def by (auto intro:continuous_intros) then show ?thesis unfolding isCont_def using filterlim_at_split by blast qed then have "\ filterlim g' at_top (at_right x)" "\ filterlim g' at_bot (at_right x)" using filterlim_at_top_nhds[of g' "at_right x"] filterlim_at_bot_nhds[of g' "at_right x"] by auto then show ?thesis unfolding jumpF_def by auto qed show "jumpF g' (at_left x) = 0" when "01" proof - have "(g' \ g' x) (at_left x)" proof (cases "x=1") case True have "continuous (at_left 1) g" using \path g\ unfolding path_def by (auto elim:continuous_on_at_left) then have "continuous (at_left x) (\t. Im(g t))" "continuous (at_left x) (\t. Re(g t))" using continuous_Im continuous_Re True by auto moreover have "Re (g (netlimit (at_left x))) \ 0" using assms(2) by (simp add: Lim_ident_at) ultimately have "continuous (at_left x) (\t. Im (g t)/Re(g t))" by (auto intro:continuous_divide) then show ?thesis unfolding g'_def continuous_def by (simp add: Lim_ident_at) next case False have "isCont (\x. Im (g x)) x" "isCont (\x. Re (g x)) x" using \path g\ unfolding path_def by (metis False atLeastAtMost_iff at_within_Icc_at continuous_Im continuous_Re continuous_on_eq_continuous_within less_le that)+ then have "isCont g' x" using assms(2) unfolding g'_def by (auto) then show ?thesis unfolding isCont_def using filterlim_at_split by blast qed then have "\ filterlim g' at_top (at_left x)" "\ filterlim g' at_bot (at_left x)" using filterlim_at_top_nhds[of g' "at_left x"] filterlim_at_bot_nhds[of g' "at_left x"] by auto then show ?thesis unfolding jumpF_def by auto qed qed lemma jump_cong: assumes "x=y" and "eventually (\x. f x=g x) (at x)" shows "jump f x = jump g y" proof - have left:"eventually (\x. f x=g x) (at_left x)" and right:"eventually (\x. f x=g x) (at_right x)" using assms(2) eventually_at_split by blast+ from filterlim_cong[OF _ _ this(1)] filterlim_cong[OF _ _ this(2)] show ?thesis unfolding jump_def using assms(1) by fastforce qed lemma jumpF_cong: assumes "F=G" and "eventually (\x. f x=g x) F" shows "jumpF f F = jumpF g G" proof - have "\\<^sub>F r in G. f r = g r" using assms(1) assms(2) by force then show ?thesis by (simp add: assms(1) filterlim_cong jumpF_def) qed lemma jump_at_left_at_right_eq: assumes "isCont f x" and "f x \ 0" and sgnx_eq:"sgnx g (at_left x) = sgnx g (at_right x)" shows "jump (\t. f t/g t) x = 0" proof - define c where "c = sgn (f x)" then have "c\0" using \f x\0\ by (simp add: sgn_zero_iff) have f_tendsto:"(f \ f x) (at_left x)" " (f \ f x) (at_right x)" using \isCont f x\ Lim_at_imp_Lim_at_within isCont_def by blast+ have False when "(g has_sgnx - c) (at_left x)" "(g has_sgnx c) (at_right x)" proof - have "sgnx g (at_left x) = -c" using that(1) by auto moreover have "sgnx g (at_right x) = c" using that(2) by auto ultimately show False using sgnx_eq \c\0\ by force qed moreover have False when "(g has_sgnx c) (at_left x)" "(g has_sgnx - c) (at_right x)" proof - have "sgnx g (at_left x) = c" using that(1) by auto moreover have "sgnx g (at_right x) = - c" using that(2) by auto ultimately show False using sgnx_eq \c\0\ by force qed ultimately show ?thesis unfolding jump_def by (auto simp add:f_tendsto filterlim_divide_at_bot_at_top_iff[OF _ \f x \ 0\] c_def) qed lemma jumpF_pos_has_sgnx: assumes "jumpF f F > 0" shows "(f has_sgnx 1) F" proof - have "filterlim f at_top F" using assms unfolding jumpF_def by argo then have "eventually (\x. f x>0) F" using filterlim_at_top_dense[of f F] by blast then show ?thesis unfolding has_sgnx_def apply (elim eventually_mono) by auto qed lemma jumpF_neg_has_sgnx: assumes "jumpF f F < 0" shows "(f has_sgnx -1) F" proof - have "filterlim f at_bot F" using assms unfolding jumpF_def by argo then have "eventually (\x. f x<0) F" using filterlim_at_bot_dense[of f F] by blast then show ?thesis unfolding has_sgnx_def apply (elim eventually_mono) by auto qed lemma jumpF_IVT: fixes f::"real \ real" and a b::real defines "right\(\(R::real \ real \ bool). R (jumpF f (at_right a)) 0 \ (continuous (at_right a) f \ R (f a) 0))" and "left\(\(R::real \ real \ bool). R (jumpF f (at_left b)) 0 \ (continuous (at_left b) f \ R (f b) 0))" assumes "a left less \ right less \ left greater" shows "\x. a x f x =0" proof - have ?thesis when "right greater" "left less" proof - have "(f has_sgnx 1) (at_right a)" proof - have ?thesis when "jumpF f (at_right a)>0" using jumpF_pos_has_sgnx[OF that] . moreover have ?thesis when "f a > 0" "continuous (at_right a) f" proof - have "(f \ f a) (at_right a)" using that(2) by (simp add: continuous_within) then show ?thesis using tendsto_nonzero_has_sgnx[of f "f a" "at_right a"] that by auto qed ultimately show ?thesis using that(1) unfolding right_def by auto qed then obtain a' where "ay. a y < a' \ f y > 0" unfolding has_sgnx_def eventually_at_right using sgn_1_pos by auto have "(f has_sgnx - 1) (at_left b)" proof - have ?thesis when "jumpF f (at_left b)<0" using jumpF_neg_has_sgnx[OF that] . moreover have ?thesis when "f b < 0" "continuous (at_left b) f" proof - have "(f \ f b) (at_left b)" using that(2) by (simp add: continuous_within) then show ?thesis using tendsto_nonzero_has_sgnx[of f "f b" "at_left b"] that by auto qed ultimately show ?thesis using that(2) unfolding left_def by auto qed then obtain b' where "b'y. b' y < b \ f y < 0" unfolding has_sgnx_def eventually_at_left using sgn_1_neg by auto have "a' \ b'" proof (rule ccontr) assume "\ a' \ b'" then have "{a<.. {b'<.. {}" using \a \b' \a by auto then obtain c where "c\{a<..{b'<..0" "f c<0" using a'_def b'_def by auto then show False by auto qed define a0 where "a0=(a+a')/2" define b0 where "b0=(b+b')/2" have [simp]:"aa \b' \a'\b'\ by auto have "f a0>0" "f b0<0" using a'_def[rule_format,of a0] b'_def[rule_format,of b0] by auto moreover have "continuous_on {a0..b0} f" using cont \a < a0\ \b0 < b\ by (meson atLeastAtMost_subseteq_greaterThanLessThan_iff continuous_on_subset) ultimately have "\x>a0. x < b0 \ f x = 0" using IVT_strict[of 0 f a0 b0] by auto then show ?thesis using \a < a0\ \b0 < b\ by (meson lessThan_strict_subset_iff psubsetE subset_psubset_trans) qed moreover have ?thesis when "right less" "left greater" proof - have "(f has_sgnx -1) (at_right a)" proof - have ?thesis when "jumpF f (at_right a)<0" using jumpF_neg_has_sgnx[OF that] . moreover have ?thesis when "f a < 0" "continuous (at_right a) f" proof - have "(f \ f a) (at_right a)" using that(2) by (simp add: continuous_within) then show ?thesis using tendsto_nonzero_has_sgnx[of f "f a" "at_right a"] that by auto qed ultimately show ?thesis using that(1) unfolding right_def by auto qed then obtain a' where "ay. a y < a' \ f y < 0" unfolding has_sgnx_def eventually_at_right using sgn_1_neg by auto have "(f has_sgnx 1) (at_left b)" proof - have ?thesis when "jumpF f (at_left b)>0" using jumpF_pos_has_sgnx[OF that] . moreover have ?thesis when "f b > 0" "continuous (at_left b) f" proof - have "(f \ f b) (at_left b)" using that(2) by (simp add: continuous_within) then show ?thesis using tendsto_nonzero_has_sgnx[of f "f b" "at_left b"] that by auto qed ultimately show ?thesis using that(2) unfolding left_def by auto qed then obtain b' where "b'y. b' y < b \ f y > 0" unfolding has_sgnx_def eventually_at_left using sgn_1_pos by auto have "a' \ b'" proof (rule ccontr) assume "\ a' \ b'" then have "{a<.. {b'<.. {}" using \a \b' \a by auto then obtain c where "c\{a<..{b'<..0" "f c<0" using a'_def b'_def by auto then show False by auto qed define a0 where "a0=(a+a')/2" define b0 where "b0=(b+b')/2" have [simp]:"aa \b' \a'\b'\ by auto have "f a0<0" "f b0>0" using a'_def[rule_format,of a0] b'_def[rule_format,of b0] by auto moreover have "continuous_on {a0..b0} f" using cont \a < a0\ \b0 < b\ by (meson atLeastAtMost_subseteq_greaterThanLessThan_iff continuous_on_subset) ultimately have "\x>a0. x < b0 \ f x = 0" using IVT_strict[of 0 f a0 b0] by auto then show ?thesis using \a < a0\ \b0 < b\ by (meson lessThan_strict_subset_iff psubsetE subset_psubset_trans) qed ultimately show ?thesis using right_left by auto qed lemma jumpF_eventually_const: assumes "eventually (\x. f x=c) F" "F\bot" shows "jumpF f F = 0" proof - have "jumpF f F = jumpF (\_. c) F" apply (rule jumpF_cong) using assms(1) by auto also have "... = 0" using jumpF_const[OF \F\bot\] by simp finally show ?thesis . qed lemma jumpF_tan_comp: "jumpF (f o tan) (at_right x) = (if cos x = 0 then jumpF f at_bot else jumpF f (at_right (tan x)))" "jumpF (f o tan) (at_left x) = (if cos x =0 then jumpF f at_top else jumpF f (at_left (tan x)))" proof - have "filtermap (f \ tan) (at_right x) = (if cos x = 0 then filtermap f at_bot else filtermap f (at_right (tan x)))" unfolding comp_def apply (subst filtermap_filtermap[of f tan,symmetric]) using filtermap_tan_at_right_inf filtermap_tan_at_right by auto then show "jumpF (f o tan) (at_right x) = (if cos x = 0 then jumpF f at_bot else jumpF f (at_right (tan x)))" unfolding jumpF_def filterlim_def by auto next have "filtermap (f \ tan) (at_left x) = (if cos x = 0 then filtermap f at_top else filtermap f (at_left (tan x)))" unfolding comp_def apply (subst filtermap_filtermap[of f tan,symmetric]) using filtermap_tan_at_left_inf filtermap_tan_at_left by auto then show "jumpF (f o tan) (at_left x) = (if cos x = 0 then jumpF f at_top else jumpF f (at_left (tan x)))" unfolding jumpF_def filterlim_def by auto qed subsection \Finite jumpFs over an interval\ definition finite_jumpFs::"(real \ real) \ real \ real \ bool" where "finite_jumpFs f a b = finite {x. (jumpF f (at_left x) \0 \ jumpF f (at_right x) \0) \ a\x \ x\b}" lemma finite_jumpFs_linear_pos: assumes "c>0" shows "finite_jumpFs (f o (\x. c * x + b)) lb ub \ finite_jumpFs f (c * lb +b) (c * ub + b)" proof - define left where "left = (\f lb ub. {x. jumpF f (at_left x) \ 0 \ lb \ x \ x \ ub})" define right where "right = (\f lb ub. {x. jumpF f (at_right x) \ 0 \ lb \ x \ x \ ub})" define g where "g=(\x. c*x+b)" define gi where "gi = (\x. (x-b)/c)" have "finite_jumpFs (f o (\x. c * x + b)) lb ub = finite (left (f o g) lb ub \ right (f o g) lb ub)" unfolding finite_jumpFs_def apply (rule arg_cong[where f=finite]) by (auto simp add:left_def right_def g_def) also have "... = finite (gi ` (left f (g lb) (g ub) \ right f (g lb) (g ub)))" proof - have j_rw: "jumpF (f o g) (at_left x) = jumpF f (at_left (g x))" "jumpF (f o g) (at_right x) = jumpF f (at_right (g x))" for x using jumpF_linear_comp[of c f b x] \c>0\ unfolding g_def by auto then have "left (f o g) lb ub = gi ` left f (g lb) (g ub)" "right (f o g) lb ub = gi ` right f (g lb) (g ub)" unfolding left_def right_def gi_def using \c>0\ by (auto simp add:g_def field_simps) then have "left (f o g) lb ub \ right (f o g) lb ub = gi ` (left f (g lb) (g ub) \ right f (g lb) (g ub))" by auto then show ?thesis by auto qed also have "... = finite (left f (g lb) (g ub) \ right f (g lb) (g ub))" apply (rule finite_image_iff) unfolding gi_def using \c >0\ inj_on_def by fastforce also have "... = finite_jumpFs f (c * lb +b) (c * ub + b)" unfolding finite_jumpFs_def apply (rule arg_cong[where f=finite]) by (auto simp add:left_def right_def g_def) finally show ?thesis . qed lemma finite_jumpFs_consts: "finite_jumpFs (\_ . c) lb ub" unfolding finite_jumpFs_def using jumpF_const by auto lemma finite_jumpFs_combine: assumes "finite_jumpFs f a b" "finite_jumpFs f b c" shows "finite_jumpFs f a c" proof - define P where "P=(\x. jumpF f (at_left x) \ 0 \ jumpF f (at_right x) \ 0)" have "{x. P x \ a \ x \ x \ c} \ {x. P x \ a \ x \ x\b} \ {x. P x \ b \x \ x\c}" by auto moreover have "finite ({x. P x \ a \ x \ x\b} \ {x. P x \ b \x \ x\c})" using assms unfolding finite_jumpFs_def P_def by auto ultimately have "finite {x. P x \ a \ x \ x \ c}" using finite_subset by auto then show ?thesis unfolding finite_jumpFs_def P_def by auto qed lemma finite_jumpFs_subE: assumes "finite_jumpFs f a b" "a\a'" "b'\b" shows "finite_jumpFs f a' b'" using assms unfolding finite_jumpFs_def apply (elim rev_finite_subset) by auto lemma finite_Psegments_Re_imp_jumpFs: assumes "finite_Psegments (\t. Re (g t - z) = 0) a b" "continuous_on {a..b} g" shows "finite_jumpFs (\t. Im (g t - z)/Re (g t - z)) a b" using assms proof (induct rule:finite_Psegments.induct) case (emptyI a b) then show ?case unfolding finite_jumpFs_def by (auto intro:rev_finite_subset[of "{a}"]) next case (insertI_1 s a b) define f where "f=(\t. Im (g t - z) / Re (g t - z))" have "finite_jumpFs f a s" proof - have "continuous_on {a..s} g" using \continuous_on {a..b} g\ \s \ {a.. by (auto elim:continuous_on_subset) then show ?thesis using insertI_1 unfolding f_def by auto qed moreover have "finite_jumpFs f s b" proof - have "jumpF f (at_left x) =0" "jumpF f (at_right x) = 0" when "x\{s<.. 0 \ jumpF f (at_right x) \ 0) \ s \ x \ x \ b} = {x. (jumpF f (at_left x) \ 0 \ jumpF f (at_right x) \ 0) \ (x=s \ x = b)}" using \s\{a.. by force then show ?thesis unfolding finite_jumpFs_def by auto qed ultimately show ?case using finite_jumpFs_combine[of _ a s b] unfolding f_def by auto next case (insertI_2 s a b) define f where "f=(\t. Im (g t - z) / Re (g t - z))" have "finite_jumpFs f a s" proof - have "continuous_on {a..s} g" using \continuous_on {a..b} g\ \s \ {a.. by (auto elim:continuous_on_subset) then show ?thesis using insertI_2 unfolding f_def by auto qed moreover have "finite_jumpFs f s b" proof - have "jumpF f (at_left x) =0" "jumpF f (at_right x) = 0" when "x\{s<..continuous_on {a..b} g\]) using insertI_2.hyps(1) that apply auto[2] using insertI_2.hyps(3) that by blast then show "jumpF f (at_left x) =0" "jumpF f (at_right x) = 0" by (simp_all add: continuous_at_split jumpF_not_infinity) qed then have "{x. (jumpF f (at_left x) \ 0 \ jumpF f (at_right x) \ 0) \ s \ x \ x \ b} = {x. (jumpF f (at_left x) \ 0 \ jumpF f (at_right x) \ 0) \ (x=s \ x = b)}" using \s\{a.. by force then show ?thesis unfolding finite_jumpFs_def by auto qed ultimately show ?case using finite_jumpFs_combine[of _ a s b] unfolding f_def by auto qed lemma finite_ReZ_segments_imp_jumpFs: assumes "finite_ReZ_segments g z" "path g" shows "finite_jumpFs (\t. Im (g t - z)/Re (g t - z)) 0 1" using assms unfolding finite_ReZ_segments_def path_def by (rule finite_Psegments_Re_imp_jumpFs) subsection \@{term jumpF} at path ends\ definition jumpF_pathstart::"(real \ complex) \ complex \ real" where "jumpF_pathstart g z= jumpF (\t. Im(g t- z)/Re(g t - z)) (at_right 0)" definition jumpF_pathfinish::"(real \ complex) \ complex \ real" where "jumpF_pathfinish g z= jumpF (\t. Im(g t - z)/Re(g t -z)) (at_left 1)" lemma jumpF_pathstart_eq_0: assumes "path g" "Re(pathstart g)\Re z" shows "jumpF_pathstart g z = 0" unfolding jumpF_pathstart_def apply (rule jumpF_im_divide_Re_0) using assms[unfolded pathstart_def] by auto lemma jumpF_pathfinish_eq_0: assumes "path g" "Re(pathfinish g)\Re z" shows "jumpF_pathfinish g z = 0" unfolding jumpF_pathfinish_def apply (rule jumpF_im_divide_Re_0) using assms[unfolded pathfinish_def] by auto lemma shows jumpF_pathfinish_reversepath: "jumpF_pathfinish (reversepath g) z = jumpF_pathstart g z" and jumpF_pathstart_reversepath: "jumpF_pathstart (reversepath g) z = jumpF_pathfinish g z" proof - define f where "f=(\t. Im (g t - z) / Re (g t - z))" define f' where "f'=(\t. Im (reversepath g t - z) / Re (reversepath g t - z))" have "f o (\t. 1 - t) = f'" unfolding f_def f'_def comp_def reversepath_def by auto then show "jumpF_pathfinish (reversepath g) z = jumpF_pathstart g z" "jumpF_pathstart (reversepath g) z = jumpF_pathfinish g z" unfolding jumpF_pathstart_def jumpF_pathfinish_def using jumpF_linear_comp(2)[of "-1" f 1 0,simplified] jumpF_linear_comp(1)[of "-1" f 1 1,simplified] apply (fold f_def f'_def) by auto qed lemma jumpF_pathstart_joinpaths[simp]: "jumpF_pathstart (g1+++g2) z = jumpF_pathstart g1 z" proof - let ?h="(\t. Im (g1 t - z) / Re (g1 t - z))" let ?f="\t. Im ((g1 +++ g2) t - z) / Re ((g1 +++ g2) t - z)" have "jumpF_pathstart g1 z = jumpF ?h (at_right 0)" unfolding jumpF_pathstart_def by simp also have "... = jumpF (?h o (\t. 2*t)) (at_right 0)" using jumpF_linear_comp[of 2 ?h 0 0,simplified] by auto also have "... = jumpF ?f (at_right 0)" proof (rule jumpF_cong) show " \\<^sub>F x in at_right 0. (?h \ (*) 2) x =?f x" unfolding eventually_at_right apply (intro exI[where x="1/2"]) by (auto simp add:joinpaths_def) qed simp also have "... =jumpF_pathstart (g1+++g2) z" unfolding jumpF_pathstart_def by simp finally show ?thesis by simp qed lemma jumpF_pathfinish_joinpaths[simp]: "jumpF_pathfinish (g1+++g2) z = jumpF_pathfinish g2 z" proof - let ?h="(\t. Im (g2 t - z) / Re (g2 t - z))" let ?f="\t. Im ((g1 +++ g2) t - z) / Re ((g1 +++ g2) t - z)" have "jumpF_pathfinish g2 z = jumpF ?h (at_left 1)" unfolding jumpF_pathfinish_def by simp also have "... = jumpF (?h o (\t. 2*t-1)) (at_left 1)" using jumpF_linear_comp[of 2 _ "-1" 1,simplified] by auto also have "... = jumpF ?f (at_left 1)" proof (rule jumpF_cong) show " \\<^sub>F x in at_left 1. (?h \ (\t. 2 * t - 1)) x =?f x" unfolding eventually_at_left apply (intro exI[where x="1/2"]) by (auto simp add:joinpaths_def) qed simp also have "... =jumpF_pathfinish (g1+++g2) z" unfolding jumpF_pathfinish_def by simp finally show ?thesis by simp qed subsection \Cauchy index\ \\Deprecated, use "cindexE" if possible\ definition cindex::"real \ real \ (real \ real) \ int" where "cindex a b f = (\x\{x. jump f x\0 \ a x real \ (real \ real) \ real" where "cindexE a b f = (\x\{x. jumpF f (at_right x) \0 \ a\x \ xx\{x. jumpF f (at_left x) \0 \ a x\b}. jumpF f (at_left x))" definition cindexE_ubd::"(real \ real) \ real" where "cindexE_ubd f = (\x\{x. jumpF f (at_right x) \0 }. jumpF f (at_right x)) - (\x\{x. jumpF f (at_left x) \0}. jumpF f (at_left x))" lemma cindexE_empty: "cindexE a a f = 0" unfolding cindexE_def by (simp add: sum.neutral) lemma cindex_const: "cindex a b (\_. c) = 0" unfolding cindex_def apply (rule sum.neutral) by auto lemma cindex_eq_cindex_poly: "cindex a b (\x. poly q x/poly p x) = cindex_poly a b q p" proof (cases "p=0") case True then show ?thesis using cindex_const by auto next case False have "cindex_poly a b q p = (\x |jump_poly q p x \0 \ a < x \ x < b. jump_poly q p x)" unfolding cindex_poly_def apply (rule sum.mono_neutral_cong_right) using jump_poly_not_root by (auto simp add: \p\0\ poly_roots_finite) also have "... = cindex a b (\x. poly q x/poly p x)" unfolding cindex_def apply (rule sum.cong) using jump_jump_poly[of q] by auto finally show ?thesis by auto qed lemma cindex_combine: assumes finite:"finite {x. jump f x\0 \ a xs. sum (jump f) ({x. jump f x\0 \ a x s))" have ssum_union:"ssum (A \ B) = ssum A + ssum B" when "A \ B ={}" for A B proof - define C where "C={x. jump f x \ 0 \ a x A" "C \ B"] that by (simp add: inf_assoc inf_sup_aci(3) inf_sup_distrib1 sum.union_disjoint) qed have "cindex a c f = ssum ({a<.. {b} \ {b<..a \b by fastforce moreover have "cindex a b f = ssum {a<..a \b by (intro sum.cong,auto) moreover have "jump f b = ssum {b}" unfolding ssum_def using \a \b by (cases "jump f b=0",auto) moreover have "cindex b c f = ssum {b<..a \b by (intro sum.cong,auto) ultimately show ?thesis apply (subst (asm) ssum_union,simp) by (subst (asm) ssum_union,auto) qed lemma cindexE_combine: assumes finite:"finite_jumpFs f a c" and "a\b" "b\c" shows "cindexE a c f = cindexE a b f + cindexE b c f" proof - define S where "S={x. (jumpF f (at_left x) \ 0 \ jumpF f (at_right x) \ 0) \ a \ x \ x \ c}" define A0 where "A0={x. jumpF f (at_right x) \ 0 \ a \ x \ x < c}" define A1 where "A1={x. jumpF f (at_right x) \ 0 \ a \ x \ x < b}" define A2 where "A2={x. jumpF f (at_right x) \ 0 \ b \ x \ x < c}" define B0 where "B0={x. jumpF f (at_left x) \ 0 \ a < x \ x \ c}" define B1 where "B1={x. jumpF f (at_left x) \ 0 \ a < x \ x \ b}" define B2 where "B2={x. jumpF f (at_left x) \ 0 \ b < x \ x \ c}" have [simp]:"finite A1" "finite A2" "finite B1" "finite B2" proof - have "finite S" using finite unfolding finite_jumpFs_def S_def by auto moreover have "A1 \ S" "A2 \ S" "B1 \ S" "B2 \ S" unfolding A1_def A2_def B1_def B2_def S_def using \a\b\ \b\c\ by auto ultimately show "finite A1" "finite A2" "finite B1" "finite B2" by (auto elim:finite_subset) qed have "cindexE a c f = sum (\x. jumpF f (at_right x)) A0 - sum (\x. jumpF f (at_left x)) B0" unfolding cindexE_def A0_def B0_def by auto also have "... = sum (\x. jumpF f (at_right x)) (A1 \ A2) - sum (\x. jumpF f (at_left x)) (B1 \ B2)" proof - have "A0=A1\A2" unfolding A0_def A1_def A2_def using assms by auto moreover have "B0=B1\B2" unfolding B0_def B1_def B2_def using assms by auto ultimately show ?thesis by auto qed also have "... = cindexE a b f + cindexE b c f" proof - have "A1 \ A2 = {}" unfolding A1_def A2_def by auto moreover have "B1 \ B2 = {}" unfolding B1_def B2_def by auto ultimately show ?thesis unfolding cindexE_def apply (fold A1_def A2_def B1_def B2_def) by (auto simp add:sum.union_disjoint) qed finally show ?thesis . qed lemma cindex_linear_comp: assumes "c\0" shows "cindex lb ub (f o (\x. c*x+b)) = (if c>0 then cindex (c*lb+b) (c*ub+b) f else - cindex (c*ub+b) (c*lb+b) f)" proof (cases "c>0") case False then have "c<0" using \c\0\ by auto have "cindex lb ub (f o (\x. c*x+b)) = - cindex (c*ub+b) (c*lb+b) f" unfolding cindex_def apply (subst sum_negf[symmetric]) apply (intro sum.reindex_cong[of "\x. (x-b)/c"]) subgoal by (simp add: inj_on_def) subgoal using False apply (subst jump_linear_comp[OF \c\0\]) by (auto simp add:\c<0\ \c\0\ field_simps) subgoal for x apply (subst jump_linear_comp[OF \c\0\]) by (auto simp add:\c<0\ \c\0\ False field_simps) done then show ?thesis using False by auto next case True have "cindex lb ub (f o (\x. c*x+b)) = cindex (c*lb+b) (c*ub+b) f" unfolding cindex_def apply (intro sum.reindex_cong[of "\x. (x-b)/c"]) subgoal by (simp add: inj_on_def) subgoal apply (subst jump_linear_comp[OF \c\0\]) by (auto simp add: True \c\0\ field_simps) subgoal for x apply (subst jump_linear_comp[OF \c\0\]) by (auto simp add: \c\0\ True field_simps) done then show ?thesis using True by auto qed lemma cindexE_linear_comp: assumes "c\0" shows "cindexE lb ub (f o (\x. c*x+b)) = (if c>0 then cindexE (c*lb+b) (c*ub+b) f else - cindexE (c*ub+b) (c*lb+b) f)" proof - define cright where "cright = (\lb ub f. (\x | jumpF f (at_right x) \ 0 \ lb \ x \ x < ub. jumpF f (at_right x)))" define cleft where "cleft = (\lb ub f. (\x | jumpF f (at_left x) \ 0 \ lb < x \ x \ ub. jumpF f (at_left x)))" have cindexE_unfold:"cindexE lb ub f = cright lb ub f - cleft lb ub f" for lb ub f unfolding cindexE_def cright_def cleft_def by auto have ?thesis when "c<0" proof - have "cright lb ub (f \ (\x. c * x + b)) = cleft (c * ub + b) (c * lb + b) f" unfolding cright_def cleft_def apply (intro sum.reindex_cong[of "\x. (x-b)/c"]) subgoal by (simp add: inj_on_def) subgoal using that by (subst jumpF_linear_comp[OF \c\0\],auto simp add:field_simps) subgoal for x using that by (subst jumpF_linear_comp[OF \c\0\],auto simp add: field_simps) done moreover have "cleft lb ub (f \ (\x. c * x + b)) = cright (c*ub+b) (c*lb + b) f" unfolding cright_def cleft_def apply (intro sum.reindex_cong[of "\x. (x-b)/c"]) subgoal by (simp add: inj_on_def) subgoal using that by (subst jumpF_linear_comp[OF \c\0\],auto simp add:field_simps) subgoal for x using that by (subst jumpF_linear_comp[OF \c\0\],auto simp add: field_simps) done ultimately show ?thesis unfolding cindexE_unfold using that by auto qed moreover have ?thesis when "c>0" proof - have "cright lb ub (f \ (\x. c * x + b)) = cright (c * lb + b) (c * ub + b) f" unfolding cright_def cleft_def apply (intro sum.reindex_cong[of "\x. (x-b)/c"]) subgoal by (simp add: inj_on_def) subgoal using that by (subst jumpF_linear_comp[OF \c\0\],auto simp add:field_simps) subgoal for x using that by (subst jumpF_linear_comp[OF \c\0\],auto simp add: field_simps) done moreover have "cleft lb ub (f \ (\x. c * x + b)) = cleft (c*lb+b) (c*ub + b) f" unfolding cright_def cleft_def apply (intro sum.reindex_cong[of "\x. (x-b)/c"]) subgoal by (simp add: inj_on_def) subgoal using that by (subst jumpF_linear_comp[OF \c\0\],auto simp add:field_simps) subgoal for x using that by (subst jumpF_linear_comp[OF \c\0\],auto simp add: field_simps) done ultimately show ?thesis unfolding cindexE_unfold using that by auto qed ultimately show ?thesis using \c\0\ by auto qed lemma cindexE_cong: assumes "finite s" and fg_eq:"\x. \as\ \ f x = g x" shows "cindexE a b f = cindexE a b g" proof - define left where "left=(\f. (\x | jumpF f (at_left x) \ 0 \ a < x \ x \ b. jumpF f (at_left x)))" define right where "right=(\f. (\x | jumpF f (at_right x) \ 0 \ a \ x \ x < b. jumpF f (at_right x)))" have "left f = left g" proof - have "jumpF f (at_left x) = jumpF g (at_left x)" when "ab" for x proof (rule jumpF_cong) define cs where "cs \ {y\s. a y (if cs = {} then (x+a)/2 else Max cs)" have "finite cs" unfolding cs_def using assms(1) by auto have "c (\y. c y f y=g y)" proof (cases "cs={}") case True then have "\y. c y y \ s" unfolding cs_def c_def by force moreover have "c=(x+a)/2" using True unfolding c_def by auto ultimately show ?thesis using fg_eq using that by auto next case False then have "c\cs" unfolding c_def using False \finite cs\ by auto moreover have "\y. c y y \ s" proof (rule ccontr) assume "\ (\y. c < y \ y < x \ y \ s) " then obtain y' where "cs" by auto then have "y'\cs" using \c\cs\ unfolding cs_def by auto then have "y'\c" unfolding c_def using False \finite cs\ by auto then show False using \c by auto qed ultimately show ?thesis unfolding cs_def using that by (auto intro!:fg_eq) qed then show "\\<^sub>F x in at_left x. f x = g x" unfolding eventually_at_left by auto qed simp then show ?thesis unfolding left_def by (auto intro: sum.cong) qed moreover have "right f = right g" proof - have "jumpF f (at_right x) = jumpF g (at_right x)" when "a\x" "x {y\s. x y (if cs = {} then (x+b)/2 else Min cs)" have "finite cs" unfolding cs_def using assms(1) by auto have "x (\y. x y f y=g y)" proof (cases "cs={}") case True then have "\y. x y y \ s" unfolding cs_def c_def by force moreover have "c=(x+b)/2" using True unfolding c_def by auto ultimately show ?thesis using fg_eq using that by auto next case False then have "c\cs" unfolding c_def using False \finite cs\ by auto moreover have "\y. x y y \ s" proof (rule ccontr) assume "\ (\y. x < y \ y < c \ y \ s) " then obtain y' where "xs" by auto then have "y'\cs" using \c\cs\ unfolding cs_def by auto then have "y'\c" unfolding c_def using False \finite cs\ by auto then show False using \c>y'\ by auto qed ultimately show ?thesis unfolding cs_def using that by (auto intro!:fg_eq) qed then show "\\<^sub>F x in at_right x. f x = g x" unfolding eventually_at_right by auto qed simp then show ?thesis unfolding right_def by (auto intro: sum.cong) qed ultimately show ?thesis unfolding cindexE_def left_def right_def by presburger qed lemma cindexE_constI: assumes "\t. \a \ f t=c" shows "cindexE a b f = 0" proof - define left where "left=(\f. (\x | jumpF f (at_left x) \ 0 \ a < x \ x \ b. jumpF f (at_left x)))" define right where "right=(\f. (\x | jumpF f (at_right x) \ 0 \ a \ x \ x < b. jumpF f (at_right x)))" have "left f = 0" proof - have "jumpF f (at_left x) = 0" when "ab" for x apply (rule jumpF_eventually_const[of _ c]) unfolding eventually_at_left using assms that by auto then show ?thesis unfolding left_def by auto qed moreover have "right f = 0" proof - have "jumpF f (at_right x) = 0" when "a\x" "x real" defines "h \ (\x. f x/g x)" assumes "ag x=0) \ a\x\x\b}" and g_imp_f:"\x\{a..b}. g x=0 \ f x\0" and f_cont:"continuous_on {a..b} f" and g_cont:"continuous_on {a..b} g" shows "cindexE a b h = jumpF h (at_right a) + cindex a b h - jumpF h (at_left b)" proof - define R where "R=(\S.{x. jumpF h (at_right x) \ 0 \ x\S})" define L where "L=(\S.{x. jumpF h (at_left x) \ 0 \ x\S})" define right where "right = (\S. (\x\R S. jumpF h (at_right x)))" define left where "left = (\S. (\x\L S. jumpF h (at_left x)))" have jump_gnz:"jumpF h (at_left x) = 0" "jumpF h (at_right x) = 0" "jump h x=0" when "a0" for x proof - have "isCont h x" unfolding h_def using f_cont g_cont that by (auto intro!:continuous_intros elim:continuous_on_interior) then show "jumpF h (at_left x) = 0" "jumpF h (at_right x) = 0" "jump h x=0" using jumpF_not_infinity jump_not_infinity unfolding continuous_at_split by auto qed have finite_jFs:"finite_jumpFs h a b" proof - define S where "S=(\s. {x. (jumpF h (at_left x) \ 0 \ jumpF h (at_right x) \ 0) \ x\s})" note jump_gnz then have "S {a<.. {x. (f x=0\g x=0) \ a\x\x\b}" unfolding S_def by auto then have "finite (S {a<.. S {a,b}" unfolding S_def using \a by auto ultimately have "finite (S {a..b})" by auto then show ?thesis unfolding S_def finite_jumpFs_def by auto qed have "cindexE a b h = right {a.. R {a..a unfolding R_def by auto moreover have "R {a..x. jumpF h (at_right x)"] unfolding right_def by simp qed moreover have "left {a<..b} = jumpF h (at_left b) + left {a<.. L {a<..b}" using False \a unfolding L_def by auto moreover have "L {a<..b} - {b} = L {a<..x. jumpF h (at_left x)"] unfolding left_def by simp qed ultimately show ?thesis by simp qed also have "... = jumpF h (at_right a) + cindex a b h - jumpF h (at_left b)" proof - define S where "S={x. g x=0 \ a < x \ x < b}" have "right {a<..x. jumpF h (at_right x)) S" unfolding right_def S_def R_def apply (rule sum.mono_neutral_left) subgoal using finite_fg by (auto elim:rev_finite_subset) subgoal using jump_gnz by auto subgoal by auto done moreover have "left {a<..x. jumpF h (at_left x)) S" unfolding left_def S_def L_def apply (rule sum.mono_neutral_left) subgoal using finite_fg by (auto elim:rev_finite_subset) subgoal using jump_gnz by auto subgoal by auto done ultimately have "right {a<..x. jumpF h (at_right x) - jumpF h (at_left x)) S" by (simp add: sum_subtractf) also have "... = sum (\x. of_int(jump h x)) S" proof (rule sum.cong) fix x assume "x\S" define hr where "hr = sgnx h (at_right x)" define hl where "hl = sgnx h (at_left x)" have "h sgnx_able (at_left x)" "hr\0" "h sgnx_able (at_right x)" "hl\0" proof - have "finite {t. h t = 0 \ a < t \ t < b}" using finite_fg unfolding h_def by (auto elim!:rev_finite_subset) moreover have "continuous_on ({a<.. a < x \ x < b}) h" unfolding h_def using f_cont g_cont by (auto intro!: continuous_intros elim:continuous_on_subset) moreover have "finite {x. g x = 0 \ a < x \ x < b}" using finite_fg by (auto elim!:rev_finite_subset) moreover have " x \ {a<..x\S\ unfolding S_def by auto ultimately show "h sgnx_able (at_left x)" "hl\0" "h sgnx_able (at_right x)" "hr\0" using finite_sgnx_at_left_at_right[of h a b "{x. g x=0 \ ax h) x" proof - have "f x\0" using \x\S\ g_imp_f unfolding S_def by auto then show ?thesis using f_cont g_cont \x\S\ unfolding h_def S_def by (auto simp add:comp_def intro!:continuous_intros elim:continuous_on_interior) qed ultimately show "jumpF h (at_right x) - jumpF h (at_left x) = real_of_int (jump h x)" using jump_jumpF[of x h] \hr\0\ \hl\0\ by auto qed auto also have "... = cindex a b h" unfolding cindex_def of_int_sum S_def apply (rule sum.mono_neutral_cong_right) using jump_gnz finite_fg by (auto elim:rev_finite_subset) finally show ?thesis by simp qed finally show ?thesis . qed subsection \Cauchy index along a path\ \\Deprecated, use "cindex\_pathE" if possible\ definition cindex_path::"(real \ complex) \ complex \ int" where "cindex_path g z = cindex 0 1 (\t. Im (g t - z) / Re (g t - z))" definition cindex_pathE::"(real \ complex) \ complex \ real" where "cindex_pathE g z = cindexE 0 1 (\t. Im (g t - z) / Re (g t - z))" lemma cindex_pathE_point: "cindex_pathE (linepath a a) b = 0" unfolding cindex_pathE_def by (simp add:cindexE_constI) lemma cindex_path_reversepath: "cindex_path (reversepath g) z = - cindex_path g z" proof - define f where "f=(\t. Im (g t - z) / Re (g t - z))" define f' where "f'=(\t. Im (reversepath g t - z) / Re (reversepath g t - z))" have "f o (\t. 1 - t) = f'" unfolding f_def f'_def comp_def reversepath_def by auto then have "cindex 0 1 f' = - cindex 0 1 f" using cindex_linear_comp[of "-1" 0 1 f 1,simplified] by simp then show ?thesis unfolding cindex_path_def apply (fold f_def f'_def) by simp qed lemma cindex_pathE_reversepath: "cindex_pathE (reversepath g) z = -cindex_pathE g z" using cindexE_linear_comp[of "-1" 0 1 "\t. (Im (g t) - Im z) / (Re (g t) - Re z)" 1] by (simp add: cindex_pathE_def reversepath_def o_def) lemma cindex_pathE_reversepath': "cindex_pathE g z = -cindex_pathE (reversepath g) z" using cindexE_linear_comp[of "-1" 0 1 "\t. (Im (g t) - Im z) / (Re (g t) - Re z)" 1] by (simp add: cindex_pathE_def reversepath_def o_def) lemma cindex_pathE_joinpaths: assumes g1:"finite_ReZ_segments g1 z" and g2: "finite_ReZ_segments g2 z" and "path g1" "path g2" "pathfinish g1 = pathstart g2" shows "cindex_pathE (g1+++g2) z = cindex_pathE g1 z + cindex_pathE g2 z" proof - define f where "f = (\g (t::real). Im (g t - z) / Re (g t - z))" have "cindex_pathE (g1 +++ g2) z = cindexE 0 1 (f (g1+++g2))" unfolding cindex_pathE_def f_def by auto also have "... = cindexE 0 (1/2) (f (g1+++g2)) + cindexE (1/2) 1 (f (g1+++g2))" proof (rule cindexE_combine) show "finite_jumpFs (f (g1 +++ g2)) 0 1" unfolding f_def apply (rule finite_ReZ_segments_imp_jumpFs) subgoal using finite_ReZ_segments_joinpaths[OF g1 g2] assms(3-5) . subgoal using path_join_imp[OF \path g1\ \path g2\ \pathfinish g1=pathstart g2\] . done qed auto also have "... = cindex_pathE g1 z + cindex_pathE g2 z" proof - have "cindexE 0 (1/2) (f (g1+++g2)) = cindex_pathE g1 z" proof - have "cindexE 0 (1/2) (f (g1+++g2)) = cindexE 0 (1/2) (f g1 o ((*) 2))" apply (rule cindexE_cong) unfolding comp_def joinpaths_def f_def by auto also have "... = cindexE 0 1 (f g1)" using cindexE_linear_comp[of 2 0 "1/2" _ 0,simplified] by simp also have "... = cindex_pathE g1 z" unfolding cindex_pathE_def f_def by auto finally show ?thesis . qed moreover have "cindexE (1/2) 1 (f (g1+++g2)) = cindex_pathE g2 z" proof - have "cindexE (1/2) 1 (f (g1+++g2)) = cindexE (1/2) 1 (f g2 o (\x. 2*x - 1))" apply (rule cindexE_cong) unfolding comp_def joinpaths_def f_def by auto also have "... = cindexE 0 1 (f g2)" using cindexE_linear_comp[of 2 "1/2" 1 _ "-1",simplified] by simp also have "... = cindex_pathE g2 z" unfolding cindex_pathE_def f_def by auto finally show ?thesis . qed ultimately show ?thesis by simp qed finally show ?thesis . qed lemma cindex_pathE_constI: assumes "\t. \0 \ g t=c" shows "cindex_pathE g z = 0" unfolding cindex_pathE_def apply (rule cindexE_constI) using assms by auto lemma cindex_pathE_subpath_combine: assumes g:"finite_ReZ_segments g z"and "path g" and "0\a" "a\b" "b\c" "c\1" shows "cindex_pathE (subpath a b g) z + cindex_pathE (subpath b c g) z = cindex_pathE (subpath a c g) z" proof - define f where "f = (\t. Im (g t - z) / Re (g t - z))" have ?thesis when "a=b" proof - have "cindex_pathE (subpath a b g) z = 0" apply (rule cindex_pathE_constI) using that unfolding subpath_def by auto then show ?thesis using that by auto qed moreover have ?thesis when "b=c" proof - have "cindex_pathE (subpath b c g) z = 0" apply (rule cindex_pathE_constI) using that unfolding subpath_def by auto then show ?thesis using that by auto qed moreover have ?thesis when "a\b" "b\c" proof - have [simp]:"aa\b\ \b\c\ by auto have "cindex_pathE (subpath a b g) z = cindexE a b f" proof - have "cindex_pathE (subpath a b g) z = cindexE 0 1 (f \ (\x. (b - a) * x + a))" unfolding cindex_pathE_def f_def comp_def subpath_def by auto also have "... = cindexE a b f" using cindexE_linear_comp[of "b-a" 0 1 f a,simplified] that(1) by auto finally show ?thesis . qed moreover have "cindex_pathE (subpath b c g) z = cindexE b c f" proof - have "cindex_pathE (subpath b c g) z = cindexE 0 1 (f \ (\x. (c - b) * x + b))" unfolding cindex_pathE_def f_def comp_def subpath_def by auto also have "... = cindexE b c f" using cindexE_linear_comp[of "c-b" 0 1 f b,simplified] that(2) by auto finally show ?thesis . qed moreover have "cindex_pathE (subpath a c g) z = cindexE a c f" proof - have "cindex_pathE (subpath a c g) z = cindexE 0 1 (f \ (\x. (c - a) * x + a))" unfolding cindex_pathE_def f_def comp_def subpath_def by auto also have "... = cindexE a c f" using cindexE_linear_comp[of "c-a" 0 1 f a,simplified] \a by auto finally show ?thesis . qed moreover have "cindexE a b f + cindexE b c f = cindexE a c f " proof - have "finite_jumpFs f a c" using finite_ReZ_segments_imp_jumpFs[OF g \path g\] \0\a\ \c\1\ unfolding f_def by (elim finite_jumpFs_subE,auto) then show ?thesis using cindexE_linear_comp cindexE_combine[OF _ \a\b\ \b\c\] by auto qed ultimately show ?thesis by auto qed ultimately show ?thesis by blast qed lemma cindex_pathE_shiftpath: assumes "finite_ReZ_segments g z" "s\{0..1}" "path g" and loop:"pathfinish g = pathstart g" shows "cindex_pathE (shiftpath s g) z = cindex_pathE g z" proof - define f where "f=(\g t. Im (g (t::real) - z) / Re (g t - z))" have "cindex_pathE (shiftpath s g) z = cindexE 0 1 (f (shiftpath s g))" unfolding cindex_pathE_def f_def by simp also have "... = cindexE 0 (1-s) (f (shiftpath s g)) + cindexE (1-s) 1 (f (shiftpath s g))" proof (rule cindexE_combine) have "finite_ReZ_segments (shiftpath s g) z" using finite_ReZ_segments_shiftpah[OF assms] . from finite_ReZ_segments_imp_jumpFs[OF this] path_shiftpath[OF \path g\ loop \s\{0..1}\] show "finite_jumpFs (f (shiftpath s g)) 0 1" unfolding f_def by simp show "0 \ 1 - s" "1 - s \ 1" using \s\{0..1}\ by auto qed also have "... = cindexE 0 s (f g) + cindexE s 1 (f g)" proof - have "cindexE 0 (1-s) (f (shiftpath s g)) = cindexE s 1 (f g)" proof - have "cindexE 0 (1-s) (f (shiftpath s g)) = cindexE 0 (1-s) ((f g) o (\t. t+s))" apply (rule cindexE_cong) unfolding shiftpath_def f_def using \s\{0..1}\ by (auto simp add:algebra_simps) also have "...= cindexE s 1 (f g)" using cindexE_linear_comp[of 1 0 "1-s" "f g" s,simplified] . finally show ?thesis . qed moreover have "cindexE (1 - s) 1 (f (shiftpath s g)) = cindexE 0 s (f g)" proof - have "cindexE (1 - s) 1 (f (shiftpath s g)) = cindexE (1-s) 1 ((f g) o (\t. t+s-1))" apply (rule cindexE_cong) unfolding shiftpath_def f_def using \s\{0..1}\ by (auto simp add:algebra_simps) also have "... = cindexE 0 s (f g)" using cindexE_linear_comp[of 1 "1-s" 1 "f g" "s-1",simplified] by (simp add:algebra_simps) finally show ?thesis . qed ultimately show ?thesis by auto qed also have "... = cindexE 0 1 (f g)" proof (rule cindexE_combine[symmetric]) show "finite_jumpFs (f g) 0 1" using finite_ReZ_segments_imp_jumpFs[OF assms(1,3)] unfolding f_def by simp show "0 \ s" "s\1" using \s\{0..1}\ by auto qed also have "... = cindex_pathE g z" unfolding cindex_pathE_def f_def by simp finally show ?thesis . qed subsection \Cauchy's Index Theorem\ theorem winding_number_cindex_pathE_aux: fixes g::"real \ complex" assumes "finite_ReZ_segments g z" and "valid_path g" "z \ path_image g" and Re_ends:"Re (g 1) = Re z" "Re (g 0) = Re z" shows "2 * Re(winding_number g z) = - cindex_pathE g z" using assms proof (induct rule:finite_ReZ_segments_induct) case (sub0 g z) have "winding_number (subpath 0 0 g) z = 0" using \z \ path_image (subpath 0 0 g)\ unfolding subpath_refl by (auto intro!: winding_number_trivial) moreover have "cindex_pathE (subpath 0 0 g) z = 0" unfolding subpath_def by (auto intro:cindex_pathE_constI) ultimately show ?case by auto next case (subEq s g z) have Re_winding_0:"Re(winding_number h z) = 0" when Re_const:"\t\{0..1}. Re (h t) = Re z" and "valid_path h" "z\path_image h" for h proof - have "Re (winding_number (\t. h t - z) 0) = (Im (Ln (pathfinish (\t. h t - z))) - Im (Ln (pathstart (\t. h t - z)))) / (2 * pi)" apply (rule Re_winding_number_half_right[of _ 0,simplified]) using Re_const \valid_path h\ \z \ path_image h\ apply auto by (metis (no_types, hide_lams) add.commute imageE le_add_same_cancel1 order_refl path_image_def plus_complex.simps(1)) moreover have "Im (Ln (h 1 - z)) = Im (Ln (h 0 - z))" proof - define z0 where "z0 = h 0 - z" define z1 where "z1 = h 1 - z" have [simp]: "z0\0" "z1\0" "Re z0=0" "Re z1=0" using \z \ path_image h\ that(1) unfolding z1_def z0_def path_image_def by auto have ?thesis when [simp]: "Im z0>0" "Im z1>0" apply (fold z1_def z0_def) using Im_Ln_eq_pi_half[of z1] Im_Ln_eq_pi_half[of z0] by auto moreover have ?thesis when [simp]: "Im z0<0" "Im z1<0" apply (fold z1_def z0_def) using Im_Ln_eq_pi_half[of z1] Im_Ln_eq_pi_half[of z0] by auto moreover have False when "Im z0\0" "Im z1\0" proof - define f where "f=(\t. Im (h t - z))" have "\x\0. x \ 1 \ f x = 0" apply (rule IVT2'[of f 1 0 0]) using that valid_path_imp_path[OF \valid_path h\] unfolding f_def z0_def z1_def path_def by (auto intro:continuous_intros) then show False using Re_const \z \ path_image h\ unfolding f_def by (metis atLeastAtMost_iff complex_surj image_eqI minus_complex.simps(2) path_defs(4) right_minus_eq) qed moreover have False when "Im z0\0" "Im z1\0" proof - define f where "f=(\t. Im (h t - z))" have "\x\0. x \ 1 \ f x = 0" apply (rule IVT') using that valid_path_imp_path[OF \valid_path h\] unfolding f_def z0_def z1_def path_def by (auto intro:continuous_intros) then show False using Re_const \z \ path_image h\ unfolding f_def by (metis atLeastAtMost_iff complex_surj image_eqI minus_complex.simps(2) path_defs(4) right_minus_eq) qed ultimately show ?thesis by argo qed ultimately have "Re (winding_number (\t. h t - z) 0) = 0" unfolding pathfinish_def pathstart_def by auto then show ?thesis using winding_number_offset by auto qed have ?case when "s = 0" proof - have *: "\t\{0..1}. Re (g t) = Re z" using \\t\{s<..<1}. Re (g t) = Re z\ \Re (g 1) = Re z\ \Re (g 0) = Re z\ \s=0\ by force have "Re(winding_number g z) = 0" by (rule Re_winding_0[OF * \valid_path g\ \z \ path_image g\]) moreover have "cindex_pathE g z = 0" unfolding cindex_pathE_def apply (rule cindexE_constI) using * by auto ultimately show ?thesis by auto qed moreover have ?case when "s\0" proof - define g1 where "g1 = subpath 0 s g" define g2 where "g2 = subpath s 1 g" have "path g" "s>0" using valid_path_imp_path[OF \valid_path g\] that \s\{0..<1}\ by auto have "2 * Re (winding_number g z) = 2*Re (winding_number g1 z) + 2*Re (winding_number g2 z)" apply (subst winding_number_subpath_combine[OF \path g\ \z\path_image g\,of 0 s 1 ,simplified,symmetric]) using \s\{0..<1}\ unfolding g1_def g2_def by auto also have "... = - cindex_pathE g1 z - cindex_pathE g2 z" proof - have "2*Re (winding_number g1 z) = - cindex_pathE g1 z" unfolding g1_def apply (rule subEq.hyps(5)) subgoal using subEq.hyps(1) subEq.prems(1) valid_path_subpath by fastforce subgoal by (meson Path_Connected.path_image_subpath_subset atLeastAtMost_iff atLeastLessThan_iff less_eq_real_def subEq(7) subEq.hyps(1) subEq.prems(1) subsetCE valid_path_imp_path zero_le_one) subgoal by (metis Groups.add_ac(2) add_0_left diff_zero mult.right_neutral subEq(2) subEq(9) subpath_def) subgoal by (simp add: subEq.prems(4) subpath_def) done moreover have "2*Re (winding_number g2 z) = - cindex_pathE g2 z" proof - have *: "\t\{0..1}. Re (g2 t) = Re z" proof fix t::real assume "t\{0..1}" have "Re (g2 t) = Re z" when "t=0 \ t=1" using that unfolding g2_def by (metis \s \ 0\ add.left_neutral diff_add_cancel mult.commute mult.left_neutral mult_zero_left subEq.hyps(2) subEq.prems(3) subpath_def) moreover have "Re (g2 t) = Re z" when "t\{0<..<1}" proof - define t' where "t'=(1 - s) * t + s" then have "t'\{s<..<1}" using that \s\{0..<1}\ unfolding t'_def apply auto by (sos "((((A<0 * (A<1 * A<2)) * R<1) + ((A<=1 * (A<0 * R<1)) * (R<1 * [1]^2))))") then have "Re (g t') = Re z" using \\t\{s<..<1}. Re (g t) = Re z\ by auto then show ?thesis unfolding g2_def subpath_def t'_def . qed ultimately show "Re (g2 t) = Re z" using \t\{0..1}\ by fastforce qed have "Re(winding_number g2 z) = 0" apply (rule Re_winding_0[OF *]) subgoal using g2_def subEq.hyps(1) subEq.prems(1) valid_path_subpath by fastforce subgoal by (metis (no_types, hide_lams) Path_Connected.path_image_subpath_subset atLeastAtMost_iff atLeastLessThan_iff g2_def less_eq_real_def subEq.hyps(1) subEq.prems(1) subEq.prems(2) subsetCE valid_path_imp_path zero_le_one) done moreover have "cindex_pathE g2 z = 0" unfolding cindex_pathE_def apply (rule cindexE_constI) using * by auto ultimately show ?thesis by auto qed ultimately show ?thesis by auto qed also have "... = - cindex_pathE g z" proof - have "finite_ReZ_segments g z" unfolding finite_ReZ_segments_def apply (rule finite_Psegments.insertI_1[of s]) subgoal using \s \ {0..<1}\ by auto subgoal using \s = 0 \ Re (g s) = Re z\ by auto subgoal using \\t\{s<..<1}. Re (g t) = Re z\ by auto subgoal proof - have "finite_Psegments (\t. Re (g (s * t)) = Re z) 0 1" using \finite_ReZ_segments (subpath 0 s g) z\ unfolding subpath_def finite_ReZ_segments_def by auto from finite_Psegments_pos_linear[of _ "1/s" 0 0 s,simplified,OF this] show "finite_Psegments (\t. Re (g t - z) = 0) 0 s" using \s>0\ unfolding comp_def by auto qed done then show ?thesis using cindex_pathE_subpath_combine[OF _ \path g\,of z 0 s 1,folded g1_def g2_def,simplified] \s\{0..<1}\ by auto qed finally show ?thesis . qed ultimately show ?case by auto next case (subNEq s g z) have Re_winding:"2*Re(winding_number h z) = jumpF_pathfinish h z - jumpF_pathstart h z" when Re_neq:"\t\{0<..<1}. Re (h t) \ Re z" and "Re (h 0) = Re z" "Re (h 1) = Re z" and "valid_path h" "z\path_image h" for h proof - have Re_winding_pos: "2*Re(winding_number h0 0) = jumpF_pathfinish h0 0 - jumpF_pathstart h0 0" when Re_gt:"\t\{0<..<1}. Re (h0 t) > 0" and "Re (h0 0) = 0" "Re (h0 1) = 0" and "valid_path h0" "0\path_image h0" for h0 proof - define f where "f \ (\(t::real). Im(h0 t) / Re (h0 t))" define ln0 where "ln0 = Im (Ln (h0 0)) / pi" define ln1 where "ln1 = Im (Ln (h0 1)) / pi" have "path h0" using \valid_path h0\ valid_path_imp_path by auto have "h0 0\0" "h0 1\0" using path_defs(4) that(5) by fastforce+ have "ln1 = jumpF_pathfinish h0 0" proof - have sgnx_at_left:"((\x. Re (h0 x)) has_sgnx 1) (at_left 1)" unfolding has_sgnx_def eventually_at_left using \\p\{0<..<1}. Re (h0 p) > 0\ by (intro exI[where x=0],auto) have cont:"continuous (at_left 1) (\t. Im (h0 t))" "continuous (at_left 1) (\t. Re (h0 t))" using \path h0\ unfolding path_def by (auto intro:continuous_on_at_left[of 0 1] continuous_intros) have ?thesis when "Im (h0 1) > 0" proof - have "ln1 = 1/2" using Im_Ln_eq_pi_half[OF \h0 1\0\] that \Re (h0 1) = 0\ unfolding ln1_def by auto moreover have "jumpF_pathfinish h0 0 = 1/2" proof - have "filterlim f at_top (at_left 1)" unfolding f_def apply (subst filterlim_divide_at_bot_at_top_iff[of _ " Im (h0 1)"]) using \Re(h0 1) = 0\ sgnx_at_left cont that unfolding continuous_within by auto then show ?thesis unfolding jumpF_pathfinish_def jumpF_def f_def by auto qed ultimately show ?thesis by auto qed moreover have ?thesis when "Im (h0 1) < 0" proof - have "ln1 = - 1/2" using Im_Ln_eq_pi_half[OF \h0 1\0\] that \Re (h0 1) = 0\ unfolding ln1_def by auto moreover have "jumpF_pathfinish h0 0 = - 1/2" proof - have "((\x. Re (h0 x)) has_sgnx - sgn (Im (h0 1))) (at_left 1)" using sgnx_at_left that by auto then have "filterlim f at_bot (at_left 1)" unfolding f_def using cont that apply (subst filterlim_divide_at_bot_at_top_iff[of _ " Im (h0 1)"]) unfolding continuous_within using \Re(h0 1) = 0\ by auto then show ?thesis unfolding jumpF_pathfinish_def jumpF_def f_def by auto qed ultimately show ?thesis by auto qed moreover have "Im (h0 1)\0" using \h0 1\0\ \Re (h0 1) = 0\ using complex.expand by auto ultimately show ?thesis by linarith qed moreover have "ln0 = jumpF_pathstart h0 0" proof - have sgnx_at_right:"((\x. Re (h0 x)) has_sgnx 1) (at_right 0)" unfolding has_sgnx_def eventually_at_right using \\p\{0<..<1}. Re (h0 p) > 0\ by (intro exI[where x=1],auto) have cont:"continuous (at_right 0) (\t. Im (h0 t))" "continuous (at_right 0) (\t. Re (h0 t))" using \path h0\ unfolding path_def by (auto intro:continuous_on_at_right[of 0 1] continuous_intros) have ?thesis when "Im (h0 0) > 0" proof - have "ln0 = 1/2" using Im_Ln_eq_pi_half[OF \h0 0\0\] that \Re (h0 0) = 0\ unfolding ln0_def by auto moreover have "jumpF_pathstart h0 0 = 1/2" proof - have "filterlim f at_top (at_right 0)" unfolding f_def apply (subst filterlim_divide_at_bot_at_top_iff[of _ " Im (h0 0)"]) using \Re(h0 0) = 0\ sgnx_at_right cont that unfolding continuous_within by auto then show ?thesis unfolding jumpF_pathstart_def jumpF_def f_def by auto qed ultimately show ?thesis by auto qed moreover have ?thesis when "Im (h0 0) < 0" proof - have "ln0 = - 1/2" using Im_Ln_eq_pi_half[OF \h0 0\0\] that \Re (h0 0) = 0\ unfolding ln0_def by auto moreover have "jumpF_pathstart h0 0 = - 1/2" proof - have "filterlim f at_bot (at_right 0)" unfolding f_def apply (subst filterlim_divide_at_bot_at_top_iff[of _ " Im (h0 0)"]) using \Re(h0 0) = 0\ sgnx_at_right cont that unfolding continuous_within by auto then show ?thesis unfolding jumpF_pathstart_def jumpF_def f_def by auto qed ultimately show ?thesis by auto qed moreover have "Im (h0 0)\0" using \h0 0\0\ \Re (h0 0) = 0\ using complex.expand by auto ultimately show ?thesis by linarith qed moreover have "2*Re(winding_number h0 0) = ln1 - ln0" proof - have "\p\path_image h0. 0 \ Re p" proof fix p assume "p \ path_image h0" then obtain t where t:"t\{0..1}" "p = h0 t" unfolding path_image_def by auto have "0 \ Re p" when "t=0 \ t=1" using that t \Re (h0 0) = 0\ \Re (h0 1) = 0\ by auto moreover have "0 \ Re p" when "t\{0<..<1}" using that t Re_gt[rule_format, of t] by fastforce ultimately show "0 \ Re p" using t(1) by fastforce qed from Re_winding_number_half_right[of _ 0,simplified,OF this \valid_path h0\ \0 \ path_image h0\] show ?thesis unfolding ln1_def ln0_def pathfinish_def pathstart_def by (auto simp add:field_simps) qed ultimately show ?thesis by auto qed have ?thesis when "\t\{0<..<1}. Re (h t) < Re z" proof - let ?hu= "\t. z - h t" have "2*Re(winding_number ?hu 0) = jumpF_pathfinish ?hu 0 - jumpF_pathstart ?hu 0" apply(rule Re_winding_pos) subgoal using that by auto subgoal using \Re (h 0) = Re z\ by auto subgoal using \Re (h 1) = Re z\ by auto subgoal using \valid_path h\ valid_path_offset valid_path_uminus_comp unfolding comp_def by fastforce subgoal using \z\path_image h\ by (simp add: image_iff path_defs(4)) done moreover have "winding_number ?hu 0 = winding_number h z" using winding_number_offset[of h z] winding_number_uminus_comp[of "\t. h t- z" 0,unfolded comp_def,simplified] \valid_path h\ \z\path_image h\ by auto moreover have "jumpF_pathfinish ?hu 0 = jumpF_pathfinish h z" unfolding jumpF_pathfinish_def apply (auto intro!:jumpF_cong eventuallyI) by (auto simp add:divide_simps algebra_simps) moreover have "jumpF_pathstart ?hu 0 = jumpF_pathstart h z" unfolding jumpF_pathstart_def apply (auto intro!:jumpF_cong eventuallyI) by (auto simp add:divide_simps algebra_simps) ultimately show ?thesis by auto qed moreover have ?thesis when "\t\{0<..<1}. Re (h t) > Re z" proof - let ?hu= "\t. h t - z" have "2*Re(winding_number ?hu 0) = jumpF_pathfinish ?hu 0 - jumpF_pathstart ?hu 0" apply(rule Re_winding_pos) subgoal using that by auto subgoal using \Re (h 0) = Re z\ by auto subgoal using \Re (h 1) = Re z\ by auto subgoal using \valid_path h\ valid_path_offset valid_path_uminus_comp unfolding comp_def by fastforce subgoal using \z\path_image h\ by simp done moreover have "winding_number ?hu 0 = winding_number h z" using winding_number_offset[of h z] \valid_path h\ \z\path_image h\ by auto moreover have "jumpF_pathfinish ?hu 0 = jumpF_pathfinish h z" unfolding jumpF_pathfinish_def by auto moreover have "jumpF_pathstart ?hu 0 = jumpF_pathstart h z" unfolding jumpF_pathstart_def by auto ultimately show ?thesis by auto qed moreover have "(\t\{0<..<1}. Re (h t) > Re z) \ (\t\{0<..<1}. Re (h t) < Re z)" proof (rule ccontr) assume " \ ((\t\{0<..<1}. Re z < Re (h t)) \ (\t\{0<..<1}. Re (h t) < Re z))" then obtain t1 t2 where t:"t1\{0<..<1}" "t2\{0<..<1}" "Re (h t1)\Re z" "Re (h t2)\Re z" unfolding path_image_def by auto have False when "t1\t2" proof - have "continuous_on {t1..t2} (\t. Re (h t))" using valid_path_imp_path[OF \valid_path h\] t unfolding path_def by (metis (full_types) atLeastatMost_subset_iff continuous_on_Re continuous_on_subset eucl_less_le_not_le greaterThanLessThan_iff) then obtain t' where t':"t'\t1" "t'\t2" "Re (h t') = Re z" using IVT'[of "\t. Re (h t)" t1 _ t2] t \t1\t2\ by auto then have "t'\{0<..<1}" using t by auto then have "Re (h t') \ Re z" using Re_neq by auto then show False using \Re (h t') = Re z\ by simp qed moreover have False when "t1\t2" proof - have "continuous_on {t2..t1} (\t. Re (h t))" using valid_path_imp_path[OF \valid_path h\] t unfolding path_def by (metis (full_types) atLeastatMost_subset_iff continuous_on_Re continuous_on_subset eucl_less_le_not_le greaterThanLessThan_iff) then obtain t' where t':"t'\t1" "t'\t2" "Re (h t') = Re z" using IVT2'[of "\t. Re (h t)" t1 _ t2] t \t1\t2\ by auto then have "t'\{0<..<1}" using t by auto then have "Re (h t') \ Re z" using Re_neq by auto then show False using \Re (h t') = Re z\ by simp qed ultimately show False by linarith qed ultimately show ?thesis by blast qed have index_ends:"cindex_pathE h z = jumpF_pathstart h z - jumpF_pathfinish h z" when Re_neq:"\t\{0<..<1}. Re (h t) \ Re z" and "valid_path h" for h proof - define f where "f = (\t. Im (h t - z) / Re (h t - z))" define Ri where "Ri = {x. jumpF f (at_right x) \ 0 \ 0 \ x \ x < 1}" define Le where "Le = {x. jumpF f (at_left x) \ 0 \ 0 < x \ x \ 1}" have "path h" using \valid_path h\ valid_path_imp_path by auto have jumpF_eq0: "jumpF f (at_left x) = 0" "jumpF f (at_right x) = 0" when "x\{0<..<1}" for x proof - have "Re (h x) \ Re z" using \\t\{0<..<1}. Re (h t) \ Re z\ that by blast then have "isCont f x" unfolding f_def using continuous_on_interior[OF \path h\[unfolded path_def]] that by (auto intro!: continuous_intros isCont_Im isCont_Re) then show "jumpF f (at_left x) = 0" "jumpF f (at_right x) = 0" unfolding continuous_at_split by (auto intro: jumpF_not_infinity) qed have "cindex_pathE h z = cindexE 0 1 f" unfolding cindex_pathE_def f_def by simp also have "... = sum (\x. jumpF f (at_right x)) Ri - sum (\x. jumpF f (at_left x)) Le" unfolding cindexE_def Ri_def Le_def by auto also have "... = jumpF f (at_right 0) - jumpF f (at_left 1)" proof - have "sum (\x. jumpF f (at_right x)) Ri = jumpF f (at_right 0)" proof (cases "jumpF f (at_right 0) = 0") case True hence False if "x \ Ri" for x using that by (cases "x = 0") (auto simp: jumpF_eq0 Ri_def) hence "Ri = {}" by blast then show ?thesis using True by auto next case False hence "x \ Ri \ x = 0" for x using that by (cases "x = 0") (auto simp: jumpF_eq0 Ri_def) hence "Ri = {0}" by blast then show ?thesis by auto qed moreover have "sum (\x. jumpF f (at_left x)) Le = jumpF f (at_left 1)" proof (cases "jumpF f (at_left 1) = 0") case True then have "Le = {}" unfolding Le_def using jumpF_eq0(1) greaterThanLessThan_iff by fastforce then show ?thesis using True by auto next case False then have "Le = {1}" unfolding Le_def using jumpF_eq0(1) greaterThanLessThan_iff by fastforce then show ?thesis by auto qed ultimately show ?thesis by auto qed also have "... = jumpF_pathstart h z - jumpF_pathfinish h z" unfolding jumpF_pathstart_def jumpF_pathfinish_def f_def by simp finally show ?thesis . qed have ?case when "s=0" proof - have "2 * Re (winding_number g z) = jumpF_pathfinish g z - jumpF_pathstart g z" apply (rule Re_winding) using subNEq that by auto moreover have "cindex_pathE g z = jumpF_pathstart g z - jumpF_pathfinish g z" apply (rule index_ends) using subNEq that by auto ultimately show ?thesis by auto qed moreover have ?case when "s\0" proof - define g1 where "g1 = subpath 0 s g" define g2 where "g2 = subpath s 1 g" have "path g" "s>0" using valid_path_imp_path[OF \valid_path g\] that \s\{0..<1}\ by auto have "2 * Re (winding_number g z) = 2*Re (winding_number g1 z) + 2*Re (winding_number g2 z)" apply (subst winding_number_subpath_combine[OF \path g\ \z\path_image g\,of 0 s 1 ,simplified,symmetric]) using \s\{0..<1}\ unfolding g1_def g2_def by auto also have "... = - cindex_pathE g1 z - cindex_pathE g2 z" proof - have "2*Re (winding_number g1 z) = - cindex_pathE g1 z" unfolding g1_def apply (rule subNEq.hyps(5)) subgoal using subNEq.hyps(1) subNEq.prems(1) valid_path_subpath by fastforce subgoal by (meson Path_Connected.path_image_subpath_subset atLeastAtMost_iff atLeastLessThan_iff less_eq_real_def subNEq(7) subNEq.hyps(1) subNEq.prems(1) subsetCE valid_path_imp_path zero_le_one) subgoal by (metis Groups.add_ac(2) add_0_left diff_zero mult.right_neutral subNEq(2) subNEq(9) subpath_def) subgoal by (simp add: subNEq.prems(4) subpath_def) done moreover have "2*Re (winding_number g2 z) = - cindex_pathE g2 z" proof - have *:"\t\{0<..<1}. Re (g2 t) \ Re z" proof fix t::real assume "t \ {0<..<1}" define t' where "t'=(1 - s) * t + s" have "t'\{s<..<1}" unfolding t'_def using \s\{0..<1}\ \t \ {0<..<1}\ apply (auto simp add:algebra_simps) by (sos "((((A<0 * (A<1 * A<2)) * R<1) + ((A<=1 * (A<1 * R<1)) * (R<1 * [1]^2))))") then have "Re (g t') \ Re z" using \\t\{s<..<1}. Re (g t) \ Re z\ by auto then show "Re (g2 t) \ Re z" unfolding g2_def subpath_def t'_def by auto qed have "2*Re (winding_number g2 z) = jumpF_pathfinish g2 z - jumpF_pathstart g2 z" apply (rule Re_winding[OF *]) subgoal by (metis add.commute add.right_neutral g2_def mult_zero_right subNEq.hyps(2) subpath_def that) subgoal by (simp add: \g2 \ subpath s 1 g\ subNEq.prems(3) subpath_def) subgoal using g2_def subNEq.hyps(1) subNEq.prems(1) valid_path_subpath by fastforce subgoal by (metis (no_types, hide_lams) Path_Connected.path_image_subpath_subset \path g\ atLeastAtMost_iff atLeastLessThan_iff g2_def less_eq_real_def subNEq.hyps(1) subNEq.prems(2) subsetCE zero_le_one) done moreover have "cindex_pathE g2 z = jumpF_pathstart g2 z - jumpF_pathfinish g2 z" apply (rule index_ends[OF *]) using g2_def subNEq.hyps(1) subNEq.prems(1) valid_path_subpath by fastforce ultimately show ?thesis by auto qed ultimately show ?thesis by auto qed also have "... = - cindex_pathE g z" proof - have "finite_ReZ_segments g z" unfolding finite_ReZ_segments_def apply (rule finite_Psegments.insertI_2[of s]) subgoal using \s \ {0..<1}\ by auto subgoal using \s = 0 \ Re (g s) = Re z\ by auto subgoal using \\t\{s<..<1}. Re (g t) \ Re z\ by auto subgoal proof - have "finite_Psegments (\t. Re (g (s * t)) = Re z) 0 1" using \finite_ReZ_segments (subpath 0 s g) z\ unfolding subpath_def finite_ReZ_segments_def by auto from finite_Psegments_pos_linear[of _ "1/s" 0 0 s,simplified,OF this] show "finite_Psegments (\t. Re (g t - z) = 0) 0 s" using \s>0\ unfolding comp_def by auto qed done then show ?thesis using cindex_pathE_subpath_combine[OF _ \path g\,of z 0 s 1,folded g1_def g2_def,simplified] \s\{0..<1}\ by auto qed finally show ?thesis . qed ultimately show ?case by auto qed theorem winding_number_cindex_pathE: fixes g::"real \ complex" assumes "finite_ReZ_segments g z" and "valid_path g" "z \ path_image g" and loop: "pathfinish g = pathstart g" shows "winding_number g z = - cindex_pathE g z / 2" proof (rule finite_ReZ_segment_cases[OF assms(1)]) fix s assume "s \ {0..<1}" "s = 0 \ Re (g s) = Re z" and const:"\t\{s<..<1}. Re (g t) = Re z" and finite:"finite_ReZ_segments (subpath 0 s g) z" have "Re (g 1) = Re z" apply(rule continuous_constant_on_closure[of "{s<..<1}" "\t. Re(g t)"]) subgoal using valid_path_imp_path[OF \valid_path g\,unfolded path_def] \s\{0..<1}\ by (auto intro!:continuous_intros continuous_Re elim:continuous_on_subset) subgoal using const by auto subgoal using \s\{0..<1}\ by auto done moreover then have "Re (g 0) = Re z" using loop unfolding path_defs by auto ultimately have "2 * Re (winding_number g z) = - cindex_pathE g z" using winding_number_cindex_pathE_aux[of g z] assms(1-3) by auto moreover have "winding_number g z \ \" using integer_winding_number[OF _ loop \z\path_image g\] valid_path_imp_path[OF \valid_path g\] by auto ultimately show "winding_number g z = - cindex_pathE g z / 2" by (metis add.right_neutral complex_eq complex_is_Int_iff mult_zero_right nonzero_mult_div_cancel_left of_real_0 zero_neq_numeral) next fix s assume "s \ {0..<1}" "s = 0 \ Re (g s) = Re z" and Re_neq:"\t\{s<..<1}. Re (g t) \ Re z" and finite:"finite_ReZ_segments (subpath 0 s g) z" have "path g" using \valid_path g\ valid_path_imp_path by auto let ?goal = "2 * Re (winding_number g z) = - cindex_pathE g z" have ?goal when "s=0" proof - have index_ends:"cindex_pathE h z = jumpF_pathstart h z - jumpF_pathfinish h z" when Re_neq:"\t\{0<..<1}. Re (h t) \ Re z" and "valid_path h" for h proof - define f where "f = (\t. Im (h t - z) / Re (h t - z))" define Ri where "Ri = {x. jumpF f (at_right x) \ 0 \ 0 \ x \ x < 1}" define Le where "Le = {x. jumpF f (at_left x) \ 0 \ 0 < x \ x \ 1}" have "path h" using \valid_path h\ valid_path_imp_path by auto have jumpF_eq0: "jumpF f (at_left x) = 0" "jumpF f (at_right x) = 0" when "x\{0<..<1}" for x proof - have "Re (h x) \ Re z" using \\t\{0<..<1}. Re (h t) \ Re z\ that by blast then have "isCont f x" unfolding f_def using continuous_on_interior[OF \path h\[unfolded path_def]] that by (auto intro!: continuous_intros isCont_Im isCont_Re) then show "jumpF f (at_left x) = 0" "jumpF f (at_right x) = 0" unfolding continuous_at_split by (auto intro: jumpF_not_infinity) qed have "cindex_pathE h z = cindexE 0 1 f" unfolding cindex_pathE_def f_def by simp also have "... = sum (\x. jumpF f (at_right x)) Ri - sum (\x. jumpF f (at_left x)) Le" unfolding cindexE_def Ri_def Le_def by auto also have "... = jumpF f (at_right 0) - jumpF f (at_left 1)" proof - have "sum (\x. jumpF f (at_right x)) Ri = jumpF f (at_right 0)" proof (cases "jumpF f (at_right 0) = 0") case True hence False if "x \ Ri" for x using that by (cases "x = 0") (auto simp: jumpF_eq0 Ri_def) hence "Ri = {}" by blast then show ?thesis using True by auto next case False hence "x \ Ri \ x = 0" for x using that by (cases "x = 0") (auto simp: jumpF_eq0 Ri_def) then have "Ri = {0}" by blast then show ?thesis by auto qed moreover have "sum (\x. jumpF f (at_left x)) Le = jumpF f (at_left 1)" proof (cases "jumpF f (at_left 1) = 0") case True then have "Le = {}" unfolding Le_def using jumpF_eq0(1) greaterThanLessThan_iff by fastforce then show ?thesis using True by auto next case False then have "Le = {1}" unfolding Le_def using jumpF_eq0(1) greaterThanLessThan_iff by fastforce then show ?thesis by auto qed ultimately show ?thesis by auto qed also have "... = jumpF_pathstart h z - jumpF_pathfinish h z" unfolding jumpF_pathstart_def jumpF_pathfinish_def f_def by simp finally show ?thesis . qed define fI where "fI=(\t. Im (g t - z))" define fR where "fR=(\t. Re (g t - z))" have fI: "(fI \ fI 0) (at_right 0)" "(fI \ fI 1) (at_left 1)" proof - have "continuous (at_right 0) fI" apply (rule continuous_on_at_right[of _ 1]) using \path g\ unfolding fI_def path_def by (auto intro:continuous_intros) then show "(fI \ fI 0) (at_right 0)" by (simp add: continuous_within) next have "continuous (at_left 1) fI" apply (rule continuous_on_at_left[of 0]) using \path g\ unfolding fI_def path_def by (auto intro:continuous_intros) then show "(fI \ fI 1) (at_left 1)" by (simp add: continuous_within) qed have fR: "(fR \ 0) (at_right 0)" "(fR \ 0) (at_left 1)" when "Re (g 0) = Re z" proof - have "continuous (at_right 0) fR" apply (rule continuous_on_at_right[of _ 1]) using \path g\ unfolding fR_def path_def by (auto intro:continuous_intros) then show "(fR \ 0) (at_right 0)" using that unfolding fR_def by (simp add: continuous_within) next have "continuous (at_left 1) fR" apply (rule continuous_on_at_left[of 0]) using \path g\ unfolding fR_def path_def by (auto intro:continuous_intros) then show "(fR \ 0) (at_left 1)" using that loop unfolding fR_def path_defs by (simp add: continuous_within) qed have "(\t\{0<..<1}. Re (g t) > Re z) \ (\t\{0<..<1}. Re (g t) < Re z)" proof (rule ccontr) assume " \ ((\t\{0<..<1}. Re z < Re (g t)) \ (\t\{0<..<1}. Re (g t) < Re z))" then obtain t1 t2 where t:"t1\{0<..<1}" "t2\{0<..<1}" "Re (g t1)\Re z" "Re (g t2)\Re z" unfolding path_image_def by auto have False when "t1\t2" proof - have "continuous_on {t1..t2} (\t. Re (g t))" using valid_path_imp_path[OF \valid_path g\] t unfolding path_def by (metis (full_types) atLeastatMost_subset_iff continuous_on_Re continuous_on_subset eucl_less_le_not_le greaterThanLessThan_iff) then obtain t' where t':"t'\t1" "t'\t2" "Re (g t') = Re z" using IVT'[of "\t. Re (g t)" t1 _ t2] t \t1\t2\ by auto then have "t'\{0<..<1}" using t by auto then have "Re (g t') \ Re z" using Re_neq \s=0\ by auto then show False using \Re (g t') = Re z\ by simp qed moreover have False when "t1\t2" proof - have "continuous_on {t2..t1} (\t. Re (g t))" using valid_path_imp_path[OF \valid_path g\] t unfolding path_def by (metis (full_types) atLeastatMost_subset_iff continuous_on_Re continuous_on_subset eucl_less_le_not_le greaterThanLessThan_iff) then obtain t' where t':"t'\t1" "t'\t2" "Re (g t') = Re z" using IVT2'[of "\t. Re (g t)" t1 _ t2] t \t1\t2\ by auto then have "t'\{0<..<1}" using t by auto then have "Re (g t') \ Re z" using Re_neq \s=0\ by auto then show False using \Re (g t') = Re z\ by simp qed ultimately show False by linarith qed moreover have ?thesis when Re_pos:"\t\{0<..<1}. Re (g t) > Re z" proof - have "Re (winding_number g z) = 0" proof - have "\p\path_image g. Re z \ Re p" proof fix p assume "p \ path_image g" then obtain t where "0\t" "t\1" "p = g t" unfolding path_image_def by auto have "Re z \ Re (g t)" apply (rule continuous_ge_on_closure[of "{0<..<1}" "\t. Re (g t)" t "Re z",simplified]) subgoal using valid_path_imp_path[OF \valid_path g\,unfolded path_def] by (auto intro:continuous_intros) subgoal using \0\t\ \t\1\ by auto subgoal for x using that[rule_format,of x] by auto done then show "Re z \ Re p" using \p = g t\ by auto qed from Re_winding_number_half_right[OF this \valid_path g\ \z\path_image g\] loop show ?thesis by auto qed moreover have "cindex_pathE g z = 0" proof - have "cindex_pathE g z = jumpF_pathstart g z - jumpF_pathfinish g z" using index_ends[OF _ \valid_path g\] Re_neq \s=0\ by auto moreover have "jumpF_pathstart g z = jumpF_pathfinish g z" when "Re (g 0) \ Re z" proof - have "jumpF_pathstart g z = 0" using jumpF_pathstart_eq_0[OF \path g\] that unfolding path_defs by auto moreover have "jumpF_pathfinish g z=0" using jumpF_pathfinish_eq_0[OF \path g\] that loop unfolding path_defs by auto ultimately show ?thesis by auto qed moreover have "jumpF_pathstart g z = jumpF_pathfinish g z" when "Re (g 0) = Re z" proof - have [simp]:"(fR has_sgnx 1) (at_right 0)" unfolding fR_def has_sgnx_def eventually_at_right apply (rule exI[where x=1]) using Re_pos by auto have [simp]:"(fR has_sgnx 1) (at_left 1)" unfolding fR_def has_sgnx_def eventually_at_left apply (rule exI[where x=0]) using Re_pos by auto have "fI 0\0" proof (rule ccontr) assume "\ fI 0 \ 0" then have "g 0 =z" using \Re (g 0) = Re z\ unfolding fI_def by (simp add: complex.expand) then show False using \z \ path_image g\ unfolding path_image_def by auto qed moreover have ?thesis when "fI 0>0" proof - have "jumpF_pathstart g z = 1/2" proof - have "(LIM x at_right 0. fI x / fR x :> at_top)" apply (subst filterlim_divide_at_bot_at_top_iff[of _ "fI 0"]) using that fI fR[OF \Re (g 0) = Re z\] by simp_all then show ?thesis unfolding jumpF_pathstart_def fI_def fR_def jumpF_def by auto qed moreover have "jumpF_pathfinish g z = 1/2" proof - have "fI 1>0" using loop that unfolding path_defs fI_def by auto then have "(LIM x at_left 1. fI x / fR x :> at_top)" apply (subst filterlim_divide_at_bot_at_top_iff[of _ "fI 1"]) using that fI fR[OF \Re (g 0) = Re z\] by simp_all then show ?thesis unfolding jumpF_pathfinish_def fI_def fR_def jumpF_def by auto qed ultimately show ?thesis by simp qed moreover have ?thesis when "fI 0<0" proof - have "jumpF_pathstart g z = - 1/2" proof - have "(LIM x at_right 0. fI x / fR x :> at_bot)" apply (subst filterlim_divide_at_bot_at_top_iff[of _ "fI 0"]) using that fI fR[OF \Re (g 0) = Re z\] by simp_all then show ?thesis unfolding jumpF_pathstart_def fI_def fR_def jumpF_def by auto qed moreover have "jumpF_pathfinish g z = - 1/2" proof - have "fI 1<0" using loop that unfolding path_defs fI_def by auto then have "(LIM x at_left 1. fI x / fR x :> at_bot)" apply (subst filterlim_divide_at_bot_at_top_iff[of _ "fI 1"]) using that fI fR[OF \Re (g 0) = Re z\] by simp_all then show ?thesis unfolding jumpF_pathfinish_def fI_def fR_def jumpF_def by auto qed ultimately show ?thesis by simp qed ultimately show ?thesis by linarith qed ultimately show ?thesis by auto qed ultimately show ?thesis by auto qed moreover have ?thesis when Re_neg:"\t\{0<..<1}. Re (g t) < Re z" proof - have "Re (winding_number g z) = 0" proof - have "\p\path_image g. Re z \ Re p" proof fix p assume "p \ path_image g" then obtain t where "0\t" "t\1" "p = g t" unfolding path_image_def by auto have "Re z \ Re (g t)" apply (rule continuous_le_on_closure[of "{0<..<1}" "\t. Re (g t)" t "Re z",simplified]) subgoal using valid_path_imp_path[OF \valid_path g\,unfolded path_def] by (auto intro:continuous_intros) subgoal using \0\t\ \t\1\ by auto subgoal for x using that[rule_format,of x] by auto done then show "Re z \ Re p" using \p = g t\ by auto qed from Re_winding_number_half_left[OF this \valid_path g\ \z\path_image g\] loop show ?thesis by auto qed moreover have "cindex_pathE g z = 0" proof - have "cindex_pathE g z = jumpF_pathstart g z - jumpF_pathfinish g z" using index_ends[OF _ \valid_path g\] Re_neq \s=0\ by auto moreover have "jumpF_pathstart g z = jumpF_pathfinish g z" when "Re (g 0) \ Re z" proof - have "jumpF_pathstart g z = 0" using jumpF_pathstart_eq_0[OF \path g\] that unfolding path_defs by auto moreover have "jumpF_pathfinish g z=0" using jumpF_pathfinish_eq_0[OF \path g\] that loop unfolding path_defs by auto ultimately show ?thesis by auto qed moreover have "jumpF_pathstart g z = jumpF_pathfinish g z" when "Re (g 0) = Re z" proof - have [simp]:"(fR has_sgnx - 1) (at_right 0)" unfolding fR_def has_sgnx_def eventually_at_right apply (rule exI[where x=1]) using Re_neg by auto have [simp]:"(fR has_sgnx - 1) (at_left 1)" unfolding fR_def has_sgnx_def eventually_at_left apply (rule exI[where x=0]) using Re_neg by auto have "fI 0\0" proof (rule ccontr) assume "\ fI 0 \ 0" then have "g 0 =z" using \Re (g 0) = Re z\ unfolding fI_def by (simp add: complex.expand) then show False using \z \ path_image g\ unfolding path_image_def by auto qed moreover have ?thesis when "fI 0>0" proof - have "jumpF_pathstart g z = - 1/2" proof - have "(LIM x at_right 0. fI x / fR x :> at_bot)" apply (subst filterlim_divide_at_bot_at_top_iff[of _ "fI 0"]) using that fI fR[OF \Re (g 0) = Re z\] by simp_all then show ?thesis unfolding jumpF_pathstart_def fI_def fR_def jumpF_def by auto qed moreover have "jumpF_pathfinish g z = - 1/2" proof - have "fI 1>0" using loop that unfolding path_defs fI_def by auto then have "(LIM x at_left 1. fI x / fR x :> at_bot)" apply (subst filterlim_divide_at_bot_at_top_iff[of _ "fI 1"]) using that fI fR[OF \Re (g 0) = Re z\] by simp_all then show ?thesis unfolding jumpF_pathfinish_def fI_def fR_def jumpF_def by auto qed ultimately show ?thesis by simp qed moreover have ?thesis when "fI 0<0" proof - have "jumpF_pathstart g z = 1/2" proof - have "(LIM x at_right 0. fI x / fR x :> at_top)" apply (subst filterlim_divide_at_bot_at_top_iff[of _ "fI 0"]) using that fI fR[OF \Re (g 0) = Re z\] by simp_all then show ?thesis unfolding jumpF_pathstart_def fI_def fR_def jumpF_def by auto qed moreover have "jumpF_pathfinish g z = 1/2" proof - have "fI 1<0" using loop that unfolding path_defs fI_def by auto then have "(LIM x at_left 1. fI x / fR x :> at_top)" apply (subst filterlim_divide_at_bot_at_top_iff[of _ "fI 1"]) using that fI fR[OF \Re (g 0) = Re z\] by simp_all then show ?thesis unfolding jumpF_pathfinish_def fI_def fR_def jumpF_def by auto qed ultimately show ?thesis by simp qed ultimately show ?thesis by linarith qed ultimately show ?thesis by auto qed ultimately show ?thesis by auto qed ultimately show ?thesis by auto qed moreover have ?goal when "s\0" proof - have "Re (g s) = Re z" using \s = 0 \ Re (g s) = Re z\ that by auto define g' where "g' = shiftpath s g" have "2 * Re (winding_number g' z) = - cindex_pathE g' z" proof (rule winding_number_cindex_pathE_aux) show "Re (g' 1) = Re z" "Re (g' 0) = Re z" using \Re (g s) = Re z\ \s\{0..<1}\ \s\0\ unfolding g'_def shiftpath_def by simp_all show "valid_path g'" using valid_path_shiftpath[OF \valid_path g\ loop,of s,folded g'_def] \s\{0..<1}\ by auto show "z \ path_image g'" using \s \ {0..<1}\ assms(3) g'_def loop path_image_shiftpath by fastforce show "finite_ReZ_segments g' z" using finite_ReZ_segments_shiftpah[OF \finite_ReZ_segments g z\ _ \path g\ loop] \s\{0..<1}\ unfolding g'_def by auto qed moreover have "winding_number g' z = winding_number g z" unfolding g'_def apply (rule winding_number_shiftpath[OF \path g\ \z \ path_image g\ loop]) using \s\{0..<1}\ by auto moreover have "cindex_pathE g' z = cindex_pathE g z" unfolding g'_def apply (rule cindex_pathE_shiftpath[OF \finite_ReZ_segments g z\ _ \path g\ loop]) using \s\{0..<1}\ by auto ultimately show ?thesis by auto qed ultimately have ?goal by auto moreover have "winding_number g z \ \" using integer_winding_number[OF _ loop \z\path_image g\] valid_path_imp_path[OF \valid_path g\] by auto ultimately show "winding_number g z = - cindex_pathE g z / 2" by (metis add.right_neutral complex_eq complex_is_Int_iff mult_zero_right nonzero_mult_div_cancel_left of_real_0 zero_neq_numeral) qed text \REMARK: The usual statement of Cauchy's Index theorem (i.e. Analytic Theory of Polynomials (2002): Theorem 11.1.3) is about the equality between the number of polynomial roots and the Cauchy index, which is the joint application of @{thm winding_number_cindex_pathE} and @{thm argument_principle}.\ end diff --git a/thys/Winding_Number_Eval/Missing_Topology.thy b/thys/Winding_Number_Eval/Missing_Topology.thy --- a/thys/Winding_Number_Eval/Missing_Topology.thy +++ b/thys/Winding_Number_Eval/Missing_Topology.thy @@ -1,740 +1,740 @@ (* Author: Wenda Li *) section \Some useful lemmas in topology\ -theory Missing_Topology imports "HOL-Analysis.Ordered_Euclidean_Space" +theory Missing_Topology imports "HOL-Analysis.Multivariate_Analysis" begin subsection \Misc\ lemma open_times_image: fixes S::"'a::real_normed_field set" assumes "open S" "c\0" shows "open (((*) c) ` S)" proof - let ?f = "\x. x/c" and ?g="((*) c)" have "continuous_on UNIV ?f" using \c\0\ by (auto intro:continuous_intros) then have "open (?f -` S)" using \open S\ by (auto elim:open_vimage) moreover have "?g ` S = ?f -` S" using \c\0\ apply auto using image_iff by fastforce ultimately show ?thesis by auto qed lemma image_linear_greaterThan: fixes x::"'a::linordered_field" assumes "c\0" shows "((\x. c*x+b) ` {x<..}) = (if c>0 then {c*x+b <..} else {..< c*x+b})" using \c\0\ apply (auto simp add:image_iff field_simps) subgoal for y by (rule bexI[where x="(y-b)/c"],auto simp add:field_simps) subgoal for y by (rule bexI[where x="(y-b)/c"],auto simp add:field_simps) done lemma image_linear_lessThan: fixes x::"'a::linordered_field" assumes "c\0" shows "((\x. c*x+b) ` {..0 then {..c\0\ apply (auto simp add:image_iff field_simps) subgoal for y by (rule bexI[where x="(y-b)/c"],auto simp add:field_simps) subgoal for y by (rule bexI[where x="(y-b)/c"],auto simp add:field_simps) done lemma continuous_on_neq_split: fixes f :: "'a::linear_continuum_topology \ 'b::linorder_topology" assumes "\x\s. f x\y" "continuous_on s f" "connected s" shows "(\x\s. f x>y) \ (\x\s. f x f ` s" using assms(1) by blast then have "(aa \ s \ y < f aa) \ aaa \ s \ f aaa < y" by (meson Topological_Spaces.connected_continuous_image assms(2) assms(3) connectedD_interval image_eqI linorder_not_le) } then show ?thesis by blast qed lemma fixes f::"'a::linorder_topology \ 'b::topological_space" assumes "continuous_on {a..b} f" "aclosure ({..a} - {b})" using \a by auto then show ?thesis using at_within_eq_bot_iff by auto qed then have "(f \ f b) (at b within {..a})" by auto moreover have "(f \ f b) (at b within {a..b})" using assms unfolding continuous_on by auto moreover have "{..a} \ {a..b} = {..b}" using \a by auto ultimately have "(f \ f b) (at b within {..b})" using Lim_Un[of f "f b" b "{..a}" "{a..b}"] by presburger then have "(f \ f b) (at b within {..closure ({b..} - {a})" using \a by auto then show ?thesis using at_within_eq_bot_iff by auto qed then have "(f \ f a) (at a within {b..})" by auto moreover have "(f \ f a) (at a within {a..b})" using assms unfolding continuous_on by auto moreover have "{b..} \ {a..b} = {a..}" using \a by auto ultimately have "(f \ f a) (at a within {a..})" using Lim_Un[of f "f a" a "{b..}" "{a..b}"] by presburger then have "(f \ f a) (at a within {a<..})" apply (elim tendsto_within_subset) by auto then show "continuous (at_right a) f" using continuous_within by auto qed subsection \More about @{term eventually}\ lemma eventually_comp_filtermap: "eventually (P o f) F \ eventually P (filtermap f F)" unfolding comp_def using eventually_filtermap by auto lemma eventually_uminus_at_top_at_bot: fixes P::"'a::{ordered_ab_group_add,linorder} \ bool" shows "eventually (P o uminus) at_bot \ eventually P at_top" "eventually (P o uminus) at_top \ eventually P at_bot" unfolding eventually_comp_filtermap by (fold at_top_mirror at_bot_mirror,auto) lemma eventually_at_infinityI: fixes P::"'a::real_normed_vector \ bool" assumes "\x. c \ norm x \ P x" shows "eventually P at_infinity" unfolding eventually_at_infinity using assms by auto lemma eventually_at_bot_linorderI: fixes c::"'a::linorder" assumes "\x. x \ c \ P x" shows "eventually P at_bot" using assms by (auto simp: eventually_at_bot_linorder) lemma eventually_times_inverse_1: fixes f::"'a \ 'b::{field,t2_space}" assumes "(f \ c) F" "c\0" shows "\\<^sub>F x in F. inverse (f x) * f x = 1" proof - have "\\<^sub>F x in F. f x\0" using assms tendsto_imp_eventually_ne by blast then show ?thesis apply (elim eventually_mono) by auto qed subsection \More about @{term filtermap}\ lemma filtermap_linear_at_within: assumes "bij f" and cont: "isCont f a" and open_map: "\S. open S \ open (f`S)" shows "filtermap f (at a within S) = at (f a) within f`S" unfolding filter_eq_iff proof safe fix P assume "eventually P (filtermap f (at a within S))" then obtain T where "open T" "a \ T" and impP:"\x\T. x\a \ x\S\ P (f x)" by (auto simp: eventually_filtermap eventually_at_topological) then show "eventually P (at (f a) within f ` S)" unfolding eventually_at_topological apply (intro exI[of _ "f`T"]) using \bij f\ open_map by (metis bij_pointE imageE imageI) next fix P assume "eventually P (at (f a) within f ` S)" then obtain T1 where "open T1" "f a \ T1" and impP:"\x\T1. x\f a \ x\f`S\ P (x)" unfolding eventually_at_topological by auto then obtain T2 where "open T2" "a \ T2" "(\x'\T2. f x' \ T1)" using cont[unfolded continuous_at_open,rule_format,of T1] by blast then have "\x\T2. x\a \ x\S\ P (f x)" using impP by (metis assms(1) bij_pointE imageI) then show "eventually P (filtermap f (at a within S))" unfolding eventually_filtermap eventually_at_topological apply (intro exI[of _ T2]) using \open T2\ \a \ T2\ by auto qed lemma filtermap_at_bot_linear_eq: fixes c::"'a::linordered_field" assumes "c\0" shows "filtermap (\x. x * c + b) at_bot = (if c>0 then at_bot else at_top)" proof (cases "c>0") case True then have "filtermap (\x. x * c + b) at_bot = at_bot" apply (intro filtermap_fun_inverse[of "\x. (x-b) / c"]) subgoal unfolding eventually_at_bot_linorder filterlim_at_bot by (auto simp add: field_simps) subgoal unfolding eventually_at_bot_linorder filterlim_at_bot by (metis mult.commute real_affinity_le) by auto then show ?thesis using \c>0\ by auto next case False then have "c<0" using \c\0\ by auto then have "filtermap (\x. x * c + b) at_bot = at_top" apply (intro filtermap_fun_inverse[of "\x. (x-b) / c"]) subgoal unfolding eventually_at_top_linorder filterlim_at_bot by (meson le_diff_eq neg_divide_le_eq) subgoal unfolding eventually_at_bot_linorder filterlim_at_top using \c < 0\ by (meson False diff_le_eq le_divide_eq) by auto then show ?thesis using \c<0\ by auto qed lemma filtermap_linear_at_left: fixes c::"'a::{linordered_field,linorder_topology,real_normed_field}" assumes "c\0" shows "filtermap (\x. c*x+b) (at_left x) = (if c>0 then at_left (c*x+b) else at_right (c*x+b))" proof - let ?f = "\x. c*x+b" have "filtermap (\x. c*x+b) (at_left x) = (at (?f x) within ?f ` {..c\0\ by (auto intro!: o_bij[of "\x. (x-b)/c"]) show "isCont ?f x" by auto show "\S. open S \ open (?f ` S)" using open_times_image[OF _ \c\0\,THEN open_translation,of _ b] by (simp add:image_image add.commute) show "at (?f x) within ?f ` {..0" using image_linear_lessThan[OF \c\0\,of b x] that by auto moreover have "?f ` {.. c>0" using image_linear_lessThan[OF \c\0\,of b x] that by auto ultimately show ?thesis by auto qed lemma filtermap_linear_at_right: fixes c::"'a::{linordered_field,linorder_topology,real_normed_field}" assumes "c\0" shows "filtermap (\x. c*x+b) (at_right x) = (if c>0 then at_right (c*x+b) else at_left (c*x+b))" proof - let ?f = "\x. c*x+b" have "filtermap ?f (at_right x) = (at (?f x) within ?f ` {x<..})" proof (subst filtermap_linear_at_within) show "bij ?f" using \c\0\ by (auto intro!: o_bij[of "\x. (x-b)/c"]) show "isCont ?f x" by auto show "\S. open S \ open (?f ` S)" using open_times_image[OF _ \c\0\,THEN open_translation,of _ b] by (simp add:image_image add.commute) show "at (?f x) within ?f ` {x<..} = at (?f x) within ?f ` {x<..}" by simp qed moreover have "?f ` {x<..} = {?f x<..}" when "c>0" using image_linear_greaterThan[OF \c\0\,of b x] that by auto moreover have "?f ` {x<..} = {.. c>0" using image_linear_greaterThan[OF \c\0\,of b x] that by auto ultimately show ?thesis by auto qed lemma filtermap_at_top_linear_eq: fixes c::"'a::linordered_field" assumes "c\0" shows "filtermap (\x. x * c + b) at_top = (if c>0 then at_top else at_bot)" proof (cases "c>0") case True then have "filtermap (\x. x * c + b) at_top = at_top" apply (intro filtermap_fun_inverse[of "\x. (x-b) / c"]) subgoal unfolding eventually_at_top_linorder filterlim_at_top by (meson le_diff_eq pos_le_divide_eq) subgoal unfolding eventually_at_top_linorder filterlim_at_top apply auto by (metis mult.commute real_le_affinity) by auto then show ?thesis using \c>0\ by auto next case False then have "c<0" using \c\0\ by auto then have "filtermap (\x. x * c + b) at_top = at_bot" apply (intro filtermap_fun_inverse[of "\x. (x-b) / c"]) subgoal unfolding eventually_at_bot_linorder filterlim_at_top by (auto simp add: field_simps) subgoal unfolding eventually_at_top_linorder filterlim_at_bot by (meson le_diff_eq neg_divide_le_eq) by auto then show ?thesis using \c<0\ by auto qed lemma filtermap_nhds_open_map: assumes cont: "isCont f a" and open_map: "\S. open S \ open (f`S)" shows "filtermap f (nhds a) = nhds (f a)" unfolding filter_eq_iff proof safe fix P assume "eventually P (filtermap f (nhds a))" then obtain S where "open S" "a \ S" "\x\S. P (f x)" by (auto simp: eventually_filtermap eventually_nhds) then show "eventually P (nhds (f a))" unfolding eventually_nhds apply (intro exI[of _ "f`S"]) by (auto intro!: open_map) qed (metis filterlim_iff tendsto_at_iff_tendsto_nhds isCont_def eventually_filtermap cont) subsection \More about @{term filterlim}\ lemma filterlim_at_infinity_times: fixes f :: "'a \ 'b::real_normed_field" assumes "filterlim f at_infinity F" "filterlim g at_infinity F" shows "filterlim (\x. f x * g x) at_infinity F" proof - have "((\x. inverse (f x) * inverse (g x)) \ 0 * 0) F" by (intro tendsto_mult tendsto_inverse assms filterlim_compose[OF tendsto_inverse_0]) then have "filterlim (\x. inverse (f x) * inverse (g x)) (at 0) F" unfolding filterlim_at using assms by (auto intro: filterlim_at_infinity_imp_eventually_ne tendsto_imp_eventually_ne eventually_conj) then show ?thesis by (subst filterlim_inverse_at_iff[symmetric]) simp_all qed lemma filterlim_at_top_at_bot[elim]: fixes f::"'a \ 'b::unbounded_dense_linorder" and F::"'a filter" assumes top:"filterlim f at_top F" and bot: "filterlim f at_bot F" and "F\bot" shows False proof - obtain c::'b where True by auto have "\\<^sub>F x in F. c < f x" using top unfolding filterlim_at_top_dense by auto moreover have "\\<^sub>F x in F. f x < c" using bot unfolding filterlim_at_bot_dense by auto ultimately have "\\<^sub>F x in F. c < f x \ f x < c" using eventually_conj by auto then have "\\<^sub>F x in F. False" by (auto elim:eventually_mono) then show False using \F\bot\ by auto qed lemma filterlim_at_top_nhds[elim]: fixes f::"'a \ 'b::{unbounded_dense_linorder,order_topology}" and F::"'a filter" assumes top:"filterlim f at_top F" and tendsto: "(f \ c) F" and "F\bot" shows False proof - obtain c'::'b where "c'>c" using gt_ex by blast have "\\<^sub>F x in F. c' < f x" using top unfolding filterlim_at_top_dense by auto moreover have "\\<^sub>F x in F. f x < c'" using order_tendstoD[OF tendsto,of c'] \c'>c\ by auto ultimately have "\\<^sub>F x in F. c' < f x \ f x < c'" using eventually_conj by auto then have "\\<^sub>F x in F. False" by (auto elim:eventually_mono) then show False using \F\bot\ by auto qed lemma filterlim_at_bot_nhds[elim]: fixes f::"'a \ 'b::{unbounded_dense_linorder,order_topology}" and F::"'a filter" assumes top:"filterlim f at_bot F" and tendsto: "(f \ c) F" and "F\bot" shows False proof - obtain c'::'b where "c'\<^sub>F x in F. c' > f x" using top unfolding filterlim_at_bot_dense by auto moreover have "\\<^sub>F x in F. f x > c'" using order_tendstoD[OF tendsto,of c'] \c' by auto ultimately have "\\<^sub>F x in F. c' < f x \ f x < c'" using eventually_conj by auto then have "\\<^sub>F x in F. False" by (auto elim:eventually_mono) then show False using \F\bot\ by auto qed lemma filterlim_at_top_linear_iff: fixes f::"'a::linordered_field \ 'b" assumes "c\0" shows "(LIM x at_top. f (x * c + b) :> F2) \ (if c>0 then (LIM x at_top. f x :> F2) else (LIM x at_bot. f x :> F2))" unfolding filterlim_def apply (subst filtermap_filtermap[of f "\x. x * c + b",symmetric]) using assms by (auto simp add:filtermap_at_top_linear_eq) lemma filterlim_at_bot_linear_iff: fixes f::"'a::linordered_field \ 'b" assumes "c\0" shows "(LIM x at_bot. f (x * c + b) :> F2) \ (if c>0 then (LIM x at_bot. f x :> F2) else (LIM x at_top. f x :> F2)) " unfolding filterlim_def apply (subst filtermap_filtermap[of f "\x. x * c + b",symmetric]) using assms by (auto simp add:filtermap_at_bot_linear_eq) lemma filterlim_tendsto_add_at_top_iff: assumes f: "(f \ c) F" shows "(LIM x F. (f x + g x :: real) :> at_top) \ (LIM x F. g x :> at_top)" proof assume "LIM x F. f x + g x :> at_top" moreover have "((\x. - f x) \ - c) F" using f by (intro tendsto_intros,simp) ultimately show "filterlim g at_top F" using filterlim_tendsto_add_at_top by fastforce qed (auto simp add:filterlim_tendsto_add_at_top[OF f]) lemma filterlim_tendsto_add_at_bot_iff: fixes c::real assumes f: "(f \ c) F" shows "(LIM x F. f x + g x :> at_bot) \ (LIM x F. g x :> at_bot)" proof - have "(LIM x F. f x + g x :> at_bot) \ (LIM x F. - f x + (- g x) :> at_top)" apply (subst filterlim_uminus_at_top) by (rule filterlim_cong,auto) also have "... = (LIM x F. - g x :> at_top)" apply (subst filterlim_tendsto_add_at_top_iff[of _ "-c"]) by (auto intro:tendsto_intros simp add:f) also have "... = (LIM x F. g x :> at_bot)" apply (subst filterlim_uminus_at_top) by (rule filterlim_cong,auto) finally show ?thesis . qed lemma tendsto_inverse_0_at_infinity: "LIM x F. f x :> at_infinity \ ((\x. inverse (f x) :: real) \ 0) F" by (metis filterlim_at filterlim_inverse_at_iff) lemma filterlim_at_infinity_divide_iff: fixes f::"'a \ 'b::real_normed_field" assumes "(f \ c) F" "c\0" shows "(LIM x F. f x / g x :> at_infinity) \ (LIM x F. g x :> at 0)" proof assume asm:"LIM x F. f x / g x :> at_infinity" have "LIM x F. inverse (f x) * (f x / g x) :> at_infinity" apply (rule tendsto_mult_filterlim_at_infinity[of _ "inverse c", OF _ _ asm]) by (auto simp add: assms(1) assms(2) tendsto_inverse) then have "LIM x F. inverse (g x) :> at_infinity" apply (elim filterlim_mono_eventually) using eventually_times_inverse_1[OF assms] by (auto elim:eventually_mono simp add:field_simps) then show "filterlim g (at 0) F" using filterlim_inverse_at_iff[symmetric] by force next assume "filterlim g (at 0) F" then have "filterlim (\x. inverse (g x)) at_infinity F" using filterlim_compose filterlim_inverse_at_infinity by blast then have "LIM x F. f x * inverse (g x) :> at_infinity" using tendsto_mult_filterlim_at_infinity[OF assms, of "\x. inverse(g x)"] by simp then show "LIM x F. f x / g x :> at_infinity" by (simp add: divide_inverse) qed lemma filterlim_tendsto_pos_mult_at_top_iff: fixes f::"'a \ real" assumes "(f \ c) F" and "0 < c" shows "(LIM x F. (f x * g x) :> at_top) \ (LIM x F. g x :> at_top)" proof assume "filterlim g at_top F" then show "LIM x F. f x * g x :> at_top" using filterlim_tendsto_pos_mult_at_top[OF assms] by auto next assume asm:"LIM x F. f x * g x :> at_top" have "((\x. inverse (f x)) \ inverse c) F" using tendsto_inverse[OF assms(1)] \0 by auto moreover have "inverse c >0" using assms(2) by auto ultimately have "LIM x F. inverse (f x) * (f x * g x) :> at_top" using filterlim_tendsto_pos_mult_at_top[OF _ _ asm,of "\x. inverse (f x)" "inverse c"] by auto then show "LIM x F. g x :> at_top" apply (elim filterlim_mono_eventually) apply simp_all[2] using eventually_times_inverse_1[OF assms(1)] \c>0\ eventually_mono by fastforce qed lemma filterlim_tendsto_pos_mult_at_bot_iff: fixes c :: real assumes "(f \ c) F" "0 < c" shows "(LIM x F. f x * g x :> at_bot) \ filterlim g at_bot F" using filterlim_tendsto_pos_mult_at_top_iff[OF assms(1,2), of "\x. - g x"] unfolding filterlim_uminus_at_bot by simp lemma filterlim_tendsto_neg_mult_at_top_iff: fixes f::"'a \ real" assumes "(f \ c) F" and "c < 0" shows "(LIM x F. (f x * g x) :> at_top) \ (LIM x F. g x :> at_bot)" proof - have "(LIM x F. f x * g x :> at_top) = (LIM x F. - g x :> at_top)" apply (rule filterlim_tendsto_pos_mult_at_top_iff[of "\x. - f x" "-c" F "\x. - g x", simplified]) using assms by (auto intro: tendsto_intros ) also have "... = (LIM x F. g x :> at_bot)" using filterlim_uminus_at_bot[symmetric] by auto finally show ?thesis . qed lemma filterlim_tendsto_neg_mult_at_bot_iff: fixes c :: real assumes "(f \ c) F" "0 > c" shows "(LIM x F. f x * g x :> at_bot) \ filterlim g at_top F" using filterlim_tendsto_neg_mult_at_top_iff[OF assms(1,2), of "\x. - g x"] unfolding filterlim_uminus_at_top by simp lemma Lim_add: fixes f g::"_ \ 'a::{t2_space,topological_monoid_add}" assumes "\y. (f \ y) F" and "\y. (g \ y) F" and "F\bot" shows "Lim F f + Lim F g = Lim F (\x. f x+g x)" apply (rule tendsto_Lim[OF \F\bot\, symmetric]) apply (auto intro!:tendsto_eq_intros) using assms tendsto_Lim by blast+ (* lemma filterlim_at_top_tendsto[elim]: fixes f::"'a \ 'b::{unbounded_dense_linorder,order_topology}" and F::"'a filter" assumes top:"filterlim f at_top F" and tendsto: "(f \ c) F" and "F\bot" shows False proof - obtain cc where "cc>c" using gt_ex by blast have "\\<^sub>F x in F. cc < f x" using top unfolding filterlim_at_top_dense by auto moreover have "\\<^sub>F x in F. f x < cc" using tendsto order_tendstoD(2)[OF _ \cc>c\] by auto ultimately have "\\<^sub>F x in F. cc < f x \ f x < cc" using eventually_conj by auto then have "\\<^sub>F x in F. False" by (auto elim:eventually_mono) then show False using \F\bot\ by auto qed lemma filterlim_at_bot_tendsto[elim]: fixes f::"'a \ 'b::{unbounded_dense_linorder,order_topology}" and F::"'a filter" assumes top:"filterlim f at_bot F" and tendsto: "(f \ c) F" and "F\bot" shows False proof - obtain cc where "cc\<^sub>F x in F. cc > f x" using top unfolding filterlim_at_bot_dense by auto moreover have "\\<^sub>F x in F. f x > cc" using tendsto order_tendstoD(1)[OF _ \cc] by auto ultimately have "\\<^sub>F x in F. cc < f x \ f x < cc" using eventually_conj by auto then have "\\<^sub>F x in F. False" by (auto elim:eventually_mono) then show False using \F\bot\ by auto qed *) subsection \Isolate and discrete\ definition (in topological_space) isolate:: "'a \ 'a set \ bool" (infixr "isolate" 60) where "x isolate S \ (x\S \ (\T. open T \ T \ S = {x}))" definition (in topological_space) discrete:: "'a set \ bool" where "discrete S \ (\x\S. x isolate S)" definition (in metric_space) uniform_discrete :: "'a set \ bool" where "uniform_discrete S \ (\e>0. \x\S. \y\S. dist x y < e \ x = y)" lemma uniformI1: assumes "e>0" "\x y. \x\S;y\S;dist x y \ x =y " shows "uniform_discrete S" unfolding uniform_discrete_def using assms by auto lemma uniformI2: assumes "e>0" "\x y. \x\S;y\S;x\y\ \ dist x y\e " shows "uniform_discrete S" unfolding uniform_discrete_def using assms not_less by blast lemma isolate_islimpt_iff:"(x isolate S) \ (\ (x islimpt S) \ x\S)" unfolding isolate_def islimpt_def by auto lemma isolate_dist_Ex_iff: fixes x::"'a::metric_space" shows "x isolate S \ (x\S \ (\e>0. \y\S. dist x y < e \ y=x))" unfolding isolate_islimpt_iff islimpt_approachable by (metis dist_commute) lemma discrete_empty[simp]: "discrete {}" unfolding discrete_def by auto lemma uniform_discrete_empty[simp]: "uniform_discrete {}" unfolding uniform_discrete_def by (simp add: gt_ex) lemma isolate_insert: fixes x :: "'a::t1_space" shows "x isolate (insert a S) \ x isolate S \ (x=a \ \ (x islimpt S))" by (meson insert_iff islimpt_insert isolate_islimpt_iff) (* TODO. Other than uniform_discrete S \ discrete S uniform_discrete S \ closed S , we should be able to prove discrete S \ closed S \ uniform_discrete S but the proof (based on Tietze Extension Theorem) seems not very trivial to me. Informal proofs can be found in http://topology.auburn.edu/tp/reprints/v30/tp30120.pdf http://msp.org/pjm/1959/9-2/pjm-v9-n2-p19-s.pdf *) lemma uniform_discrete_imp_closed: "uniform_discrete S \ closed S" by (meson discrete_imp_closed uniform_discrete_def) lemma uniform_discrete_imp_discrete: "uniform_discrete S \ discrete S" by (metis discrete_def isolate_dist_Ex_iff uniform_discrete_def) lemma isolate_subset:"x isolate S \ T \ S \ x\T \ x isolate T" unfolding isolate_def by fastforce lemma discrete_subset[elim]: "discrete S \ T \ S \ discrete T" unfolding discrete_def using islimpt_subset isolate_islimpt_iff by blast lemma uniform_discrete_subset[elim]: "uniform_discrete S \ T \ S \ uniform_discrete T" by (meson subsetD uniform_discrete_def) lemma continuous_on_discrete: "discrete S \ continuous_on S f" unfolding continuous_on_topological by (metis discrete_def islimptI isolate_islimpt_iff) (* Is euclidean_space really necessary?*) lemma uniform_discrete_insert: fixes S :: "'a::euclidean_space set" shows "uniform_discrete (insert a S) \ uniform_discrete S" proof assume asm:"uniform_discrete S" let ?thesis = "uniform_discrete (insert a S)" have ?thesis when "a\S" using that asm by (simp add: insert_absorb) moreover have ?thesis when "S={}" using that asm by (simp add: uniform_discrete_def) moreover have ?thesis when "a\S" "S\{}" proof - obtain e1 where "e1>0" and e1_dist:"\x\S. \y\S. dist y x < e1 \ y = x" using asm unfolding uniform_discrete_def by auto define e2 where "e2 \ min (setdist {a} S) e1" have "closed S" using asm uniform_discrete_imp_closed by auto then have "e2>0" by (simp add: \0 < e1\ e2_def setdist_gt_0_compact_closed that(1) that(2)) moreover have "x = y" when "x\insert a S" "y\insert a S" "dist x y < e2 " for x y proof - have ?thesis when "x=a" "y=a" using that by auto moreover have ?thesis when "x=a" "y\S" using that setdist_le_dist[of x "{a}" y S] \dist x y < e2\ unfolding e2_def by fastforce moreover have ?thesis when "y=a" "x\S" using that setdist_le_dist[of y "{a}" x S] \dist x y < e2\ unfolding e2_def by (simp add: dist_commute) moreover have ?thesis when "x\S" "y\S" using e1_dist[rule_format, OF that] \dist x y < e2\ unfolding e2_def by (simp add: dist_commute) ultimately show ?thesis using that by auto qed ultimately show ?thesis unfolding uniform_discrete_def by meson qed ultimately show ?thesis by auto qed (simp add: subset_insertI uniform_discrete_subset) lemma discrete_compact_finite_iff: fixes S :: "'a::t1_space set" shows "discrete S \ compact S \ finite S" proof assume "finite S" then have "compact S" using finite_imp_compact by auto moreover have "discrete S" unfolding discrete_def using isolate_islimpt_iff islimpt_finite[OF \finite S\] by auto ultimately show "discrete S \ compact S" by auto next assume "discrete S \ compact S" then show "finite S" by (meson discrete_def Heine_Borel_imp_Bolzano_Weierstrass isolate_islimpt_iff order_refl) qed lemma uniform_discrete_finite_iff: fixes S :: "'a::heine_borel set" shows "uniform_discrete S \ bounded S \ finite S" proof assume "uniform_discrete S \ bounded S" then have "discrete S" "compact S" using uniform_discrete_imp_discrete uniform_discrete_imp_closed compact_eq_bounded_closed by auto then show "finite S" using discrete_compact_finite_iff by auto next assume asm:"finite S" let ?thesis = "uniform_discrete S \ bounded S" have ?thesis when "S={}" using that by auto moreover have ?thesis when "S\{}" proof - have "\x. \d>0. \y\S. y \ x \ d \ dist x y" using finite_set_avoid[OF \finite S\] by auto then obtain f where f_pos:"f x>0" and f_dist: "\y\S. y \ x \ f x \ dist x y" if "x\S" for x by metis define f_min where "f_min \ Min (f ` S)" have "f_min > 0" unfolding f_min_def by (simp add: asm f_pos that) moreover have "\x\S. \y\S. f_min > dist x y \ x=y" using f_dist unfolding f_min_def by (metis Min_gr_iff all_not_in_conv asm dual_order.irrefl eq_iff finite_imageI imageI less_eq_real_def) ultimately have "uniform_discrete S" unfolding uniform_discrete_def by auto moreover have "bounded S" using \finite S\ by auto ultimately show ?thesis by auto qed ultimately show ?thesis by blast qed lemma uniform_discrete_image_scale: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "uniform_discrete S" and dist:"\x\S. \y\S. dist x y = c * dist (f x) (f y)" shows "uniform_discrete (f ` S)" proof - have ?thesis when "S={}" using that by auto moreover have ?thesis when "S\{}" "c\0" proof - obtain x1 where "x1\S" using \S\{}\ by auto have ?thesis when "S-{x1} = {}" proof - have "S={x1}" using that \S\{}\ by auto then show ?thesis using uniform_discrete_insert[of "f x1"] by auto qed moreover have ?thesis when "S-{x1} \ {}" proof - obtain x2 where "x2\S-{x1}" using \S-{x1} \ {}\ by auto then have "x2\S" "x1\x2" by auto then have "dist x1 x2 > 0" by auto moreover have "dist x1 x2 = c * dist (f x1) (f x2)" using dist[rule_format, OF \x1\S\ \x2\S\] . moreover have "dist (f x2) (f x2) \ 0" by auto ultimately have False using \c\0\ by (simp add: zero_less_mult_iff) then show ?thesis by auto qed ultimately show ?thesis by auto qed moreover have ?thesis when "S\{}" "c>0" proof - obtain e1 where "e1>0" and e1_dist:"\x\S. \y\S. dist y x < e1 \ y = x" using \uniform_discrete S\ unfolding uniform_discrete_def by auto define e where "e= e1/c" have "x1 = x2" when "x1\ f ` S" "x2\ f ` S" "dist x1 x2 < e " for x1 x2 proof - obtain y1 where y1:"y1\S" "x1=f y1" using \x1\ f ` S\ by auto obtain y2 where y2:"y2\S" "x2=f y2" using \x2\ f ` S\ by auto have "dist y1 y2 < e1" using dist[rule_format, OF y1(1) y2(1)] \c>0\ \dist x1 x2 < e\ unfolding e_def apply (fold y1(2) y2(2)) by (auto simp add:divide_simps mult.commute) then have "y1=y2" using e1_dist[rule_format, OF y2(1) y1(1)] by simp then show "x1=x2" using y1(2) y2(2) by auto qed moreover have "e>0" using \e1>0\ \c>0\ unfolding e_def by auto ultimately show ?thesis unfolding uniform_discrete_def by meson qed ultimately show ?thesis by fastforce qed end diff --git a/thys/pGCL/Misc.thy b/thys/pGCL/Misc.thy --- a/thys/pGCL/Misc.thy +++ b/thys/pGCL/Misc.thy @@ -1,455 +1,455 @@ (* * Copyright (C) 2014 NICTA * All rights reserved. *) (* Author: David Cock - David.Cock@nicta.com.au *) section \Miscellaneous Mathematics\ theory Misc imports - "HOL-Analysis.Ordered_Euclidean_Space" + "HOL-Analysis.Multivariate_Analysis" begin text_raw \\label{s:misc}\ lemma sum_UNIV: fixes S::"'a::finite set" assumes complete: "\x. x\S \ f x = 0" shows "sum f S = sum f UNIV" proof - from complete have "sum f S = sum f (UNIV - S) + sum f S" by(simp) also have "... = sum f UNIV" by(auto intro: sum.subset_diff[symmetric]) finally show ?thesis . qed lemma cInf_mono: fixes A::"'a::conditionally_complete_lattice set" assumes lower: "\b. b \ B \ \a\A. a \ b" and bounded: "\a. a \ A \ c \ a" and ne: "B \ {}" shows "Inf A \ Inf B" proof(rule cInf_greatest[OF ne]) fix b assume bin: "b \ B" with lower obtain a where ain: "a \ A" and le: "a \ b" by(auto) from ain bounded have "Inf A \ a" by(intro cInf_lower bdd_belowI, auto) also note le finally show "Inf A \ b" . qed lemma max_distrib: fixes c::real assumes nn: "0 \ c" shows "c * max a b = max (c * a) (c * b)" proof(cases "a \ b") case True moreover with nn have "c * a \ c * b" by(auto intro:mult_left_mono) ultimately show ?thesis by(simp add:max.absorb2) next case False then have "b \ a" by(auto) moreover with nn have "c * b \ c * a" by(auto intro:mult_left_mono) ultimately show ?thesis by(simp add:max.absorb1) qed lemma mult_div_mono_left: fixes c::real assumes nnc: "0 \ c" and nzc: "c \ 0" and inv: "a \ inverse c * b" shows "c * a \ b" proof - from nnc inv have "c * a \ (c * inverse c) * b" by(auto simp:mult.assoc intro:mult_left_mono) also from nzc have "... = b" by(simp) finally show "c * a \ b" . qed lemma mult_div_mono_right: fixes c::real assumes nnc: "0 \ c" and nzc: "c \ 0" and inv: "inverse c * a \ b" shows "a \ c * b" proof - from nzc have "a = (c * inverse c) * a " by(simp) also from nnc inv have "(c * inverse c) * a \ c * b" by(auto simp:mult.assoc intro:mult_left_mono) finally show "a \ c * b" . qed lemma min_distrib: fixes c::real assumes nnc: "0 \ c" shows "c * min a b = min (c * a) (c * b)" proof(cases "a \ b") case True moreover with nnc have "c * a \ c * b" by(blast intro:mult_left_mono) ultimately show ?thesis by(auto) next case False hence "b \ a" by(auto) moreover with nnc have "c * b \ c * a" by(blast intro:mult_left_mono) ultimately show ?thesis by(simp add:min.absorb2) qed lemma finite_set_least: fixes S::"'a::linorder set" assumes finite: "finite S" and ne: "S \ {}" shows "\x\S. \y\S. x \ y" proof - have "S = {} \ (\x\S. \y\S. x \ y)" proof(rule finite_induct, simp_all add:assms) fix x::'a and S::"'a set" assume IH: "S = {} \ (\x\S. \y\S. x \ y)" show "(\y\S. x \ y) \ (\x'\S. x' \ x \ (\y\S. x' \ y))" proof(cases "S={}") case True then show ?thesis by(auto) next case False with IH have "\x\S. \y\S. x \ y" by(auto) then obtain z where zin: "z \ S" and zmin: "\y\S. z \ y" by(auto) thus ?thesis by(cases "z \ x", auto) qed qed with ne show ?thesis by(auto) qed lemma cSup_add: fixes c::real assumes ne: "S \ {}" and bS: "\x. x\S \ x \ b" shows "Sup S + c = Sup {x + c |x. x \ S}" proof(rule antisym) from ne bS show "Sup {x + c |x. x \ S} \ Sup S + c" by(auto intro!:cSup_least add_right_mono cSup_upper bdd_aboveI) have "Sup S \ Sup {x + c |x. x \ S} - c" proof(intro cSup_least ne) fix x assume xin: "x \ S" from bS have "\x. x\S \ x + c \ b + c" by(auto intro:add_right_mono) hence "bdd_above {x + c |x. x \ S}" by(intro bdd_aboveI, blast) with xin have "x + c \ Sup {x + c |x. x \ S}" by(auto intro:cSup_upper) thus "x \ Sup {x + c |x. x \ S} - c" by(auto) qed thus "Sup S + c \ Sup {x + c |x. x \ S}" by(auto) qed lemma cSup_mult: fixes c::real assumes ne: "S \ {}" and bS: "\x. x\S \ x \ b" and nnc: "0 \ c" shows "c * Sup S = Sup {c * x |x. x \ S}" proof(cases) assume "c = 0" moreover from ne have "\x. x \ S" by(auto) ultimately show ?thesis by(simp) next assume cnz: "c \ 0" show ?thesis proof(rule antisym) from bS have baS: "bdd_above S" by(intro bdd_aboveI, auto) with ne nnc show "Sup {c * x |x. x \ S} \ c * Sup S" by(blast intro!:cSup_least mult_left_mono[OF cSup_upper]) have "Sup S \ inverse c * Sup {c * x |x. x \ S}" proof(intro cSup_least ne) fix x assume xin: "x\S" moreover from bS nnc have "\x. x\S \ c * x \ c * b" by(auto intro:mult_left_mono) ultimately have "c * x \ Sup {c * x |x. x \ S}" by(auto intro!:cSup_upper bdd_aboveI) moreover from nnc have "0 \ inverse c" by(auto) ultimately have "inverse c * (c * x) \ inverse c * Sup {c * x |x. x \ S}" by(auto intro:mult_left_mono) with cnz show "x \ inverse c * Sup {c * x |x. x \ S}" by(simp add:mult.assoc[symmetric]) qed with nnc have "c * Sup S \ c * (inverse c * Sup {c * x |x. x \ S})" by(auto intro:mult_left_mono) with cnz show "c * Sup S \ Sup {c * x |x. x \ S}" by(simp add:mult.assoc[symmetric]) qed qed lemma closure_contains_Sup: fixes S :: "real set" assumes neS: "S \ {}" and bS: "\x\S. x \ B" shows "Sup S \ closure S" proof - let ?T = "uminus ` S" from neS have neT: "?T \ {}" by(auto) from bS have bT: "\x\?T. -B \ x" by(auto) hence bbT: "bdd_below ?T" by(intro bdd_belowI, blast) have "Sup S = - Inf ?T" proof(rule antisym) from neT bbT have "\x. x\S \ Inf (uminus ` S) \ -x" by(blast intro:cInf_lower) hence "\x. x\S \ -1 * -x \ -1 * Inf (uminus ` S)" by(rule mult_left_mono_neg, auto) hence lenInf: "\x. x\S \ x \ - Inf (uminus ` S)" by(simp) with neS bS show "Sup S \ - Inf ?T" by(blast intro:cSup_least) have "- Sup S \ Inf ?T" proof(rule cInf_greatest[OF neT]) fix x assume "x \ uminus ` S" then obtain y where yin: "y \ S" and rwx: "x = -y" by(auto) from yin bS have "y \ Sup S" by(intro cSup_upper bdd_belowI, auto) hence "-1 * Sup S \ -1 * y" by(simp add:mult_left_mono_neg) with rwx show "- Sup S \ x" by(simp) qed hence "-1 * Inf ?T \ -1 * (- Sup S)" by(simp add:mult_left_mono_neg) thus "- Inf ?T \ Sup S" by(simp) qed also { from neT bbT have "Inf ?T \ closure ?T" by(rule closure_contains_Inf) hence "- Inf ?T \ uminus ` closure ?T" by(auto) } also { have "linear uminus" by(auto intro:linearI) hence "uminus ` closure ?T \ closure (uminus ` ?T)" by(rule closure_linear_image_subset) } also { have "uminus ` ?T \ S" by(auto) hence "closure (uminus ` ?T) \ closure S" by(rule closure_mono) } finally show "Sup S \ closure S" . qed lemma tendsto_min: fixes x y::real assumes ta: "a \ x" and tb: "b \ y" shows "(\i. min (a i) (b i)) \ min x y" proof(rule LIMSEQ_I, simp) fix e::real assume pe: "0 < e" from ta pe obtain noa where balla: "\n\noa. abs (a n - x) < e" by(auto dest:LIMSEQ_D) from tb pe obtain nob where ballb: "\n\nob. abs (b n - y) < e" by(auto dest:LIMSEQ_D) { fix n assume ge: "max noa nob \ n" hence gea: "noa \ n" and geb: "nob \ n" by(auto) have "abs (min (a n) (b n) - min x y) < e" proof cases assume le: "min (a n) (b n) \ min x y" show ?thesis proof cases assume "a n \ b n" hence rwmin: "min (a n) (b n) = a n" by(auto) with le have "a n \ min x y" by(simp) moreover from gea balla have "abs (a n - x) < e" by(simp) moreover have "min x y \ x" by(auto) ultimately have "abs (a n - min x y) < e" by(auto) with rwmin show "abs (min (a n) (b n) - min x y) < e" by(simp) next assume "\ a n \ b n" hence "b n \ a n" by(auto) hence rwmin: "min (a n) (b n) = b n" by(auto) with le have "b n \ min x y" by(simp) moreover from geb ballb have "abs (b n - y) < e" by(simp) moreover have "min x y \ y" by(auto) ultimately have "abs (b n - min x y) < e" by(auto) with rwmin show "abs (min (a n) (b n) - min x y) < e" by(simp) qed next assume "\ min (a n) (b n) \ min x y" hence le: "min x y \ min (a n) (b n)" by(auto) show ?thesis proof cases assume "x \ y" hence rwmin: "min x y = x" by(auto) with le have "x \ min (a n) (b n)" by(simp) moreover from gea balla have "abs (a n - x) < e" by(simp) moreover have "min (a n) (b n) \ a n" by(auto) ultimately have "abs (min (a n) (b n) - x) < e" by(auto) with rwmin show "abs (min (a n) (b n) - min x y) < e" by(simp) next assume "\ x \ y" hence "y \ x" by(auto) hence rwmin: "min x y = y" by(auto) with le have "y \ min (a n) (b n)" by(simp) moreover from geb ballb have "abs (b n - y) < e" by(simp) moreover have "min (a n) (b n) \ b n" by(auto) ultimately have "abs (min (a n) (b n) - y) < e" by(auto) with rwmin show "abs (min (a n) (b n) - min x y) < e" by(simp) qed qed } thus "\no. \n\no. \min (a n) (b n) - min x y\ < e" by(blast) qed definition supp :: "('s \ real) \ 's set" where "supp f = {x. f x \ 0}" definition dist_remove :: "('s \ real) \ 's \ 's \ real" where "dist_remove p x = (\y. if y=x then 0 else p y / (1 - p x))" lemma supp_dist_remove: "p x \ 0 \ p x \ 1 \ supp (dist_remove p x) = supp p - {x}" by(auto simp:dist_remove_def supp_def) lemma supp_empty: "supp f = {} \ f x = 0" by(simp add:supp_def) lemma nsupp_zero: "x \ supp f \ f x = 0" by(simp add:supp_def) lemma sum_supp: fixes f::"'a::finite \ real" shows "sum f (supp f) = sum f UNIV" proof - have "sum f (UNIV - supp f) = 0" by(simp add:supp_def) hence "sum f (supp f) = sum f (UNIV - supp f) + sum f (supp f)" by(simp) also have "... = sum f UNIV" by(simp add:sum.subset_diff[symmetric]) finally show ?thesis . qed subsection \Truncated Subtraction\ text_raw \\label{s:trunc_sub}\ definition tminus :: "real \ real \ real" (infixl "\" 60) where "x \ y = max (x - y) 0" lemma minus_le_tminus[intro!,simp]: "a - b \ a \ b" unfolding tminus_def by(auto) lemma tminus_cancel_1: "0 \ a \ a + 1 \ 1 = a" unfolding tminus_def by(simp) lemma tminus_zero_imp_le: "x \ y \ 0 \ x \ y" by(simp add:tminus_def) lemma tminus_zero[simp]: "0 \ x \ x \ 0 = x" by(simp add:tminus_def) lemma tminus_left_mono: "a \ b \ a \ c \ b \ c" unfolding tminus_def by(case_tac "a \ c", simp_all) lemma tminus_less: "\ 0 \ a; 0 \ b \ \ a \ b \ a" unfolding tminus_def by(force) lemma tminus_left_distrib: assumes nna: "0 \ a" shows "a * (b \ c) = a * b \ a * c" proof(cases "b \ c") case True note le = this hence "a * max (b - c) 0 = 0" by(simp add:max.absorb2) also { from nna le have "a * b \ a * c" by(blast intro:mult_left_mono) hence "0 = max (a * b - a * c) 0" by(simp add:max.absorb1) } finally show ?thesis by(simp add:tminus_def) next case False hence le: "c \ b" by(auto) hence "a * max (b - c) 0 = a * (b - c)" by(simp only:max.absorb1) also { from nna le have "a * c \ a * b" by(blast intro:mult_left_mono) hence "a * (b - c) = max (a * b - a * c) 0" by(simp add:max.absorb1 field_simps) } finally show ?thesis by(simp add:tminus_def) qed lemma tminus_le[simp]: "b \ a \ a \ b = a - b" unfolding tminus_def by(simp) lemma tminus_le_alt[simp]: "a \ b \ a \ b = 0" by(simp add:tminus_def) lemma tminus_nle[simp]: "\b \ a \ a \ b = 0" unfolding tminus_def by(simp) lemma tminus_add_mono: "(a+b) \ (c+d) \ (a\c) + (b\d)" proof(cases "0 \ a - c") case True note pac = this show ?thesis proof(cases "0 \ b - d") case True note pbd = this from pac and pbd have "(c + d) \ (a + b)" by(simp) with pac and pbd show ?thesis by(simp) next case False with pac show ?thesis by(cases "c + d \ a + b", auto) qed next case False note nac = this show ?thesis proof(cases "0 \ b - d") case True with nac show ?thesis by(cases "c + d \ a + b", auto) next case False note nbd = this with nac have "\(c + d) \ (a + b)" by(simp) with nac and nbd show ?thesis by(simp) qed qed lemma tminus_sum_mono: assumes fS: "finite S" shows "sum f S \ sum g S \ sum (\x. f x \ g x) S" (is "?X S") proof(rule finite_induct) from fS show "finite S" . show "?X {}" by(simp) fix x and F assume fF: "finite F" and xniF: "x \ F" and IH: "?X F" have "f x + sum f F \ g x + sum g F \ (f x \ g x) + (sum f F \ sum g F)" by(rule tminus_add_mono) also from IH have "... \ (f x \ g x) + (\x\F. f x \ g x)" by(rule add_left_mono) finally show "?X (insert x F)" by(simp add:sum.insert[OF fF xniF]) qed lemma tminus_nneg[simp,intro]: "0 \ a \ b" by(cases "b \ a", auto) lemma tminus_right_antimono: assumes clb: "c \ b" shows "a \ b \ a \ c" proof(cases "b \ a") case True moreover with clb have "c \ a" by(auto) moreover note clb ultimately show ?thesis by(simp) next case False then show ?thesis by(simp) qed lemma min_tminus_distrib: "min a b \ c = min (a \ c) (b \ c)" unfolding tminus_def by(auto) end