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.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 + e_def Let_def pdevs_val_sum 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/Intersection.thy b/thys/Affine_Arithmetic/Intersection.thy --- a/thys/Affine_Arithmetic/Intersection.thy +++ b/thys/Affine_Arithmetic/Intersection.thy @@ -1,2807 +1,2814 @@ section \Intersection\ theory Intersection imports "HOL-Library.Monad_Syntax" Polygon Counterclockwise_2D_Arbitrary Affine_Form begin text \\label{sec:intersection}\ subsection \Polygons and @{term ccw}, @{term lex}, @{term psi}, @{term coll}\ lemma polychain_of_ccw_conjunction: assumes sorted: "ccw'.sortedP 0 Ps" assumes z: "z \ set (polychain_of Pc Ps)" shows "list_all (\(xi, xj). ccw xi xj (fst z) \ ccw xi xj (snd z)) (polychain_of Pc Ps)" using assms proof (induction Ps arbitrary: Pc z rule: list.induct) case (Cons P Ps) { assume "set Ps = {}" hence ?case using Cons by simp } moreover { assume "set Ps \ {}" hence "Ps \ []" by simp { fix a assume "a \ set Ps" hence "ccw' 0 P a" using Cons.prems by (auto elim!: linorder_list0.sortedP_Cons) } note ccw' = this have sorted': "linorder_list0.sortedP (ccw' 0) Ps" using Cons.prems by (auto elim!: linorder_list0.sortedP_Cons) from in_set_polychain_of_imp_sum_list[OF Cons(3)] obtain d where d: "z = (Pc + sum_list (take d (P # Ps)), Pc + sum_list (take (Suc d) (P # Ps)))" . from Cons(3) have disj: "z = (Pc, Pc + P) \ z \ set (polychain_of (Pc + P) Ps)" by auto let ?th = "\(xi, xj). ccw xi xj Pc \ ccw xi xj (Pc + P)" have la: "list_all ?th (polychain_of (Pc + P) Ps)" proof (rule list_allI) fix x assume x: "x \ set (polychain_of (Pc + P) Ps)" from in_set_polychain_of_imp_sum_list[OF this] obtain e where e: "x = (Pc + P + sum_list (take e Ps), Pc + P + sum_list (take (Suc e) Ps))" by auto { assume "e \ length Ps" hence "?th x" by (auto simp: e) } moreover { assume "e < length Ps" have 0: "\e. e < length Ps \ ccw' 0 P (Ps ! e)" by (rule ccw') (simp add: ) have 2: "0 < e \ ccw' 0 (P + sum_list (take e Ps)) (Ps ! e)" using \e < length Ps\ by (auto intro!: ccw'.add1 0 ccw'.sum2 sorted' ccw'.sorted_nth_less simp: sum_list_sum_nth) have "ccw Pc (Pc + P + sum_list (take e Ps)) (Pc + P + sum_list (take (Suc e) Ps))" by (cases "e = 0") (auto simp add: ccw_translate_origin take_Suc_eq add.assoc[symmetric] 0 2 intro!: ccw'_imp_ccw intro: cyclic) hence "ccw (Pc + P + sum_list (take e Ps)) (Pc + P + sum_list (take (Suc e) Ps)) Pc" by (rule cyclic) moreover have "0 < e \ ccw 0 (Ps ! e) (- sum_list (take e Ps))" using \e < length Ps\ by (auto simp add: take_Suc_eq add.assoc[symmetric] sum_list_sum_nth intro!: ccw'_imp_ccw ccw'.sum2 sorted' ccw'.sorted_nth_less) hence "ccw (Pc + P + sum_list (take e Ps)) (Pc + P + sum_list (take (Suc e) Ps)) (Pc + P)" by (cases "e = 0") (simp_all add: ccw_translate_origin take_Suc_eq) ultimately have "?th x" by (auto simp add: e) } ultimately show "?th x" by arith qed from disj have ?case proof assume z: "z \ set (polychain_of (Pc + P) Ps)" have "ccw 0 P (sum_list (take d (P # Ps)))" proof (cases d) case (Suc e) note e = this show ?thesis proof (cases e) case (Suc f) have "ccw 0 P (P + sum_list (take (Suc f) Ps))" using z by (force simp add: sum_list_sum_nth intro!: ccw'.sum intro: ccw' ccw'_imp_ccw) thus ?thesis by (simp add: e Suc) qed (simp add: e) qed simp hence "ccw Pc (Pc + P) (fst z)" by (simp add: d ccw_translate_origin) moreover from z have "ccw 0 P (P + sum_list (take d Ps))" by (cases d, force) (force simp add: sum_list_sum_nth intro!: ccw'_imp_ccw ccw'.sum intro: ccw')+ hence "ccw Pc (Pc + P) (snd z)" by (simp add: d ccw_translate_origin) moreover from z Cons.prems have "list_all (\(xi, xj). ccw xi xj (fst z) \ ccw xi xj (snd z)) (polychain_of (Pc + P) Ps)" by (intro Cons.IH) (auto elim!: linorder_list0.sortedP_Cons) ultimately show ?thesis by simp qed (simp add: la) } ultimately show ?case by blast qed simp lemma lex_polychain_of_center: "d \ set (polychain_of x0 xs) \ list_all (\x. lex x 0) xs \ lex (snd d) x0" proof (induction xs arbitrary: x0) case (Cons x xs) thus ?case by (auto simp add: lex_def lex_translate_origin dest!: Cons.IH) qed (auto simp: lex_translate_origin) lemma lex_polychain_of_last: "(c, d) \ set (polychain_of x0 xs) \ list_all (\x. lex x 0) xs \ lex (snd (last (polychain_of x0 xs))) d" proof (induction xs arbitrary: x0 c d) case (Cons x xs) let ?c1 = "c = x0 \ d = x0 + x" let ?c2 = "(c, d) \ set (polychain_of (x0 + x) xs)" from Cons(2) have "?c1 \ ?c2" by auto thus ?case proof assume ?c1 with Cons.prems show ?thesis by (auto intro!: lex_polychain_of_center) next assume ?c2 from Cons.IH[OF this] Cons.prems show ?thesis by auto qed qed (auto simp: lex_translate_origin) lemma distinct_fst_polychain_of: assumes "list_all (\x. x \ 0) xs" assumes "list_all (\x. lex x 0) xs" shows "distinct (map fst (polychain_of x0 xs))" using assms proof (induction xs arbitrary: x0) case Nil thus ?case by simp next case (Cons x xs) hence "\d. list_all (\x. lex x 0) (x # take d xs)" by (auto simp: list_all_iff dest!: in_set_takeD) from sum_list_nlex_eq_zero_iff[OF this] Cons.prems show ?case by (cases "xs = []") (auto intro!: Cons.IH elim!: in_set_polychain_of_imp_sum_list) qed lemma distinct_snd_polychain_of: assumes "list_all (\x. x \ 0) xs" assumes "list_all (\x. lex x 0) xs" shows "distinct (map snd (polychain_of x0 xs))" using assms proof (induction xs arbitrary: x0) case Nil thus ?case by simp next case (Cons x xs) have contra: "\d. xs \ [] \ list_all (\x. x \ 0) xs \ list_all ((=) 0) (take (Suc d) xs) \ False" by (auto simp: neq_Nil_conv) from Cons have "\d. list_all (\x. lex x 0) (take (Suc d) xs)" by (auto simp: list_all_iff dest!: in_set_takeD) from sum_list_nlex_eq_zero_iff[OF this] Cons.prems contra show ?case by (cases "xs = []") (auto intro!: Cons.IH elim!: in_set_polychain_of_imp_sum_list dest!: contra) qed subsection \Orient all entries\ lift_definition nlex_pdevs::"point pdevs \ point pdevs" is "\x n. if lex 0 (x n) then - x n else x n" by simp lemma pdevs_apply_nlex_pdevs[simp]: "pdevs_apply (nlex_pdevs x) n = (if lex 0 (pdevs_apply x n) then - pdevs_apply x n else pdevs_apply x n)" by transfer simp lemma nlex_pdevs_zero_pdevs[simp]: "nlex_pdevs zero_pdevs = zero_pdevs" by (auto intro!: pdevs_eqI) lemma pdevs_domain_nlex_pdevs[simp]: "pdevs_domain (nlex_pdevs x) = pdevs_domain x" by (auto simp: pdevs_domain_def) lemma degree_nlex_pdevs[simp]: "degree (nlex_pdevs x) = degree x" by (rule degree_cong) auto lemma pdevs_val_nlex_pdevs: assumes "e \ UNIV \ I" "uminus ` I = I" obtains e' where "e' \ UNIV \ I" "pdevs_val e x = pdevs_val e' (nlex_pdevs x)" using assms by (atomize_elim, intro exI[where x="\n. if lex 0 (pdevs_apply x n) then - e n else e n"]) (force simp: pdevs_val_pdevs_domain intro!: sum.cong) lemma pdevs_val_nlex_pdevs2: assumes "e \ UNIV \ I" "uminus ` I = I" obtains e' where "e' \ UNIV \ I" "pdevs_val e (nlex_pdevs x) = pdevs_val e' x" using assms by (atomize_elim, intro exI[where x="\n. (if lex 0 (pdevs_apply x n) then - e n else e n)"]) (force simp: pdevs_val_pdevs_domain intro!: sum.cong) lemma pdevs_val_selsort_ccw: assumes "distinct xs" assumes "e \ UNIV \ I" obtains e' where "e' \ UNIV \ I" "pdevs_val e (pdevs_of_list xs) = pdevs_val e' (pdevs_of_list (ccw.selsort 0 xs))" proof - have "set xs = set (ccw.selsort 0 xs)" "distinct xs" "distinct (ccw.selsort 0 xs)" by (simp_all add: assms) from this assms(2) obtain e' where "e' \ UNIV \ I" "pdevs_val e (pdevs_of_list xs) = pdevs_val e' (pdevs_of_list (ccw.selsort 0 xs))" by (rule pdevs_val_permute) thus thesis .. qed lemma pdevs_val_selsort_ccw2: assumes "distinct xs" assumes "e \ UNIV \ I" obtains e' where "e' \ UNIV \ I" "pdevs_val e (pdevs_of_list (ccw.selsort 0 xs)) = pdevs_val e' (pdevs_of_list xs)" proof - have "set (ccw.selsort 0 xs) = set xs" "distinct (ccw.selsort 0 xs)" "distinct xs" by (simp_all add: assms) from this assms(2) obtain e' where "e' \ UNIV \ I" "pdevs_val e (pdevs_of_list (ccw.selsort 0 xs)) = pdevs_val e' (pdevs_of_list xs)" by (rule pdevs_val_permute) thus thesis .. qed lemma lex_nlex_pdevs: "lex (pdevs_apply (nlex_pdevs x) i) 0" by (auto simp: lex_def algebra_simps prod_eq_iff) subsection \Lowest Vertex\ definition lowest_vertex::"'a::ordered_euclidean_space aform \ 'a" where "lowest_vertex X = fst X - sum_list (map snd (list_of_pdevs (snd X)))" lemma snd_abs: "snd (abs x) = abs (snd x)" by (metis abs_prod_def snd_conv) lemma lowest_vertex: fixes X Y::"(real*real) aform" assumes "e \ UNIV \ {-1 .. 1}" assumes "\i. snd (pdevs_apply (snd X) i) \ 0" assumes "\i. abs (snd (pdevs_apply (snd Y) i)) = abs (snd (pdevs_apply (snd X) i))" assumes "degree_aform Y = degree_aform X" assumes "fst Y = fst X" shows "snd (lowest_vertex X) \ snd (aform_val e Y)" proof - from abs_pdevs_val_le_tdev[OF assms(1), of "snd Y"] have "snd \pdevs_val e (snd Y)\ \ (\isnd (pdevs_apply (snd X) i)\)" unfolding lowest_vertex_def by (auto simp: aform_val_def tdev_def less_eq_prod_def snd_sum snd_abs assms) also have "\ = (\i \ snd (sum_list (map snd (list_of_pdevs (snd X))))" by (simp add: sum_list_list_of_pdevs dense_list_of_pdevs_def sum_list_distinct_conv_sum_set snd_sum atLeast0LessThan) finally show ?thesis by (auto simp: aform_val_def lowest_vertex_def minus_le_iff snd_abs abs_real_def assms split: if_split_asm) qed lemma sum_list_nonposI: fixes xs::"'a::ordered_comm_monoid_add list" shows "list_all (\x. x \ 0) xs \ sum_list xs \ 0" by (induct xs) (auto simp: intro!: add_nonpos_nonpos) lemma center_le_lowest: "fst (fst X) \ fst (lowest_vertex (fst X, nlex_pdevs (snd X)))" by (auto simp: lowest_vertex_def fst_sum_list intro!: sum_list_nonposI) (auto simp: lex_def list_all_iff list_of_pdevs_def dest!: in_set_butlastD split: if_split_asm) lemma lowest_vertex_eq_center_iff: "lowest_vertex (x0, nlex_pdevs (snd X)) = x0 \ snd X = zero_pdevs" proof assume "lowest_vertex (x0, nlex_pdevs (snd X)) = x0" then have "sum_list (map snd (list_of_pdevs (nlex_pdevs (snd X)))) = 0" by (simp add: lowest_vertex_def) moreover have "list_all (\x. Counterclockwise_2D_Arbitrary.lex x 0) (map snd (list_of_pdevs (nlex_pdevs (snd X))))" by (auto simp add: list_all_iff list_of_pdevs_def) ultimately have "\x\set (list_of_pdevs (nlex_pdevs (snd X))). snd x = 0" by (simp add: sum_list_nlex_eq_zero_iff list_all_iff) then have "pdevs_apply (snd X) i = pdevs_apply zero_pdevs i" for i by (simp add: list_of_pdevs_def split: if_splits) then show "snd X = zero_pdevs" by (rule pdevs_eqI) qed (simp add: lowest_vertex_def) subsection \Collinear Generators\ lemma scaleR_le_self_cancel: fixes c::"'a::ordered_real_vector" shows "a *\<^sub>R c \ c \ (1 < a \ c \ 0 \ a < 1 \ 0 \ c \ a = 1)" using scaleR_le_0_iff[of "a - 1" c] by (simp add: algebra_simps) lemma pdevs_val_coll: assumes coll: "list_all (coll 0 x) xs" assumes nlex: "list_all (\x. lex x 0) xs" assumes "x \ 0" assumes "f \ UNIV \ {-1 .. 1}" obtains e where "e \ {-1 .. 1}" "pdevs_val f (pdevs_of_list xs) = e *\<^sub>R (sum_list xs)" proof cases assume "sum_list xs = 0" have "pdevs_of_list xs = zero_pdevs" by (auto intro!: pdevs_eqI sum_list_nlex_eq_zeroI[OF nlex \sum_list xs = 0\] simp: pdevs_apply_pdevs_of_list list_all_iff dest!: nth_mem) hence "0 \ {-1 .. 1::real}" "pdevs_val f (pdevs_of_list xs) = 0 *\<^sub>R sum_list xs" by simp_all thus ?thesis .. next assume "sum_list xs \ 0" have "sum_list (map abs xs) \ 0" by (auto intro!: sum_list_nonneg) hence [simp]: "\sum_list (map abs xs) \ 0" by (metis \sum_list xs \ 0\ abs_le_zero_iff antisym_conv sum_list_abs) have collist: "list_all (coll 0 (sum_list xs)) xs" proof (rule list_allI) fix y assume "y \ set xs" hence "coll 0 x y" using coll by (auto simp: list_all_iff) also have "coll 0 x (sum_list xs)" using coll by (auto simp: list_all_iff intro!: coll_sum_list) finally (coll_trans) show "coll 0 (sum_list xs) y" by (simp add: coll_commute \x \ 0\) qed { fix i assume "i < length xs" hence "\r. xs ! i = r *\<^sub>R (sum_list xs)" by (metis (mono_tags) coll_scale nth_mem \sum_list xs \ 0\ list_all_iff collist) } then obtain r where r: "\i. i < length xs \ (xs ! i) = r i *\<^sub>R (sum_list xs)" by metis let ?coll = "pdevs_of_list xs" have "pdevs_val f (pdevs_of_list xs) = (\iR xs ! i)" unfolding pdevs_val_sum by (simp add: pdevs_apply_pdevs_of_list less_degree_pdevs_of_list_imp_less_length) also have "\ = (\iR (sum_list xs))" by (simp add: r less_degree_pdevs_of_list_imp_less_length) also have "\ = (\iR (sum_list xs)" by (simp add: algebra_simps scaleR_sum_left) finally have eq: "pdevs_val f ?coll = (\iR (sum_list xs)" (is "_ = ?e *\<^sub>R _") . have "abs (pdevs_val f ?coll) \ tdev ?coll" using assms(4) by (intro abs_pdevs_val_le_tdev) (auto simp: Pi_iff less_imp_le) also have "\ = sum_list (map abs xs)" using assms by simp also note eq finally have less: "\?e\ *\<^sub>R abs (sum_list xs) \ sum_list (map abs xs)" by (simp add: abs_scaleR) also have "\sum_list xs\ = sum_list (map abs xs)" using coll \x \ 0\ nlex by (rule abs_sum_list_coll) finally have "?e \ {-1 .. 1}" by (auto simp add: less_le scaleR_le_self_cancel) thus ?thesis using eq .. qed lemma scaleR_eq_self_cancel: fixes x::"'a::real_vector" shows "a *\<^sub>R x = x \ a = 1 \ x = 0" using scaleR_cancel_right[of a x 1] by simp lemma abs_pdevs_val_less_tdev: assumes "e \ UNIV \ {-1 <..< 1}" "degree x > 0" shows "\pdevs_val e x\ < tdev x" proof - have bnds: "\i. \e i\ < 1" "\i. \e i\ \ 1" using assms by (auto simp: Pi_iff abs_less_iff order.strict_implies_order) moreover let ?w = "degree x - 1" have witness: "\e ?w\ *\<^sub>R \pdevs_apply x ?w\ < \pdevs_apply x ?w\" using degree_least_nonzero[of x] assms bnds by (intro neq_le_trans) (auto simp: scaleR_eq_self_cancel Pi_iff intro!: scaleR_left_le_one_le neq_le_trans intro: abs_leI less_imp_le dest!: order.strict_implies_not_eq) ultimately show ?thesis using assms witness bnds by (auto simp: pdevs_val_sum tdev_def abs_scaleR intro!: le_less_trans[OF sum_abs] sum_strict_mono_ex1 scaleR_left_le_one_le) qed lemma pdevs_val_coll_strict: assumes coll: "list_all (coll 0 x) xs" assumes nlex: "list_all (\x. lex x 0) xs" assumes "x \ 0" assumes "f \ UNIV \ {-1 <..< 1}" obtains e where "e \ {-1 <..< 1}" "pdevs_val f (pdevs_of_list xs) = e *\<^sub>R (sum_list xs)" proof cases assume "sum_list xs = 0" have "pdevs_of_list xs = zero_pdevs" by (auto intro!: pdevs_eqI sum_list_nlex_eq_zeroI[OF nlex \sum_list xs = 0\] simp: pdevs_apply_pdevs_of_list list_all_iff dest!: nth_mem) hence "0 \ {-1 <..< 1::real}" "pdevs_val f (pdevs_of_list xs) = 0 *\<^sub>R sum_list xs" by simp_all thus ?thesis .. next assume "sum_list xs \ 0" have "sum_list (map abs xs) \ 0" by (auto intro!: sum_list_nonneg) hence [simp]: "\sum_list (map abs xs) \ 0" by (metis \sum_list xs \ 0\ abs_le_zero_iff antisym_conv sum_list_abs) have "\x \ set xs. x \ 0" proof (rule ccontr) assume "\ (\x\set xs. x \ 0)" hence "\x. x \ set xs \ x = 0" by auto hence "sum_list xs = 0" by (auto simp: sum_list_eq_0_iff_nonpos list_all_iff less_eq_prod_def prod_eq_iff fst_sum_list snd_sum_list) thus False using \sum_list xs \ 0\ by simp qed then obtain i where i: "i < length xs" "xs ! i \ 0" by (metis in_set_conv_nth) hence "i < degree (pdevs_of_list xs)" by (auto intro!: degree_gt simp: pdevs_apply_pdevs_of_list) hence deg_pos: "0 < degree (pdevs_of_list xs)" by simp have collist: "list_all (coll 0 (sum_list xs)) xs" proof (rule list_allI) fix y assume "y \ set xs" hence "coll 0 x y" using coll by (auto simp: list_all_iff) also have "coll 0 x (sum_list xs)" using coll by (auto simp: list_all_iff intro!: coll_sum_list) finally (coll_trans) show "coll 0 (sum_list xs) y" by (simp add: coll_commute \x \ 0\) qed { fix i assume "i < length xs" hence "\r. xs ! i = r *\<^sub>R (sum_list xs)" by (metis (mono_tags, lifting) \sum_list xs \ 0\ coll_scale collist list_all_iff nth_mem) } then obtain r where r: "\i. i < length xs \ (xs ! i) = r i *\<^sub>R (sum_list xs)" by metis let ?coll = "pdevs_of_list xs" have "pdevs_val f (pdevs_of_list xs) = (\iR xs ! i)" unfolding pdevs_val_sum by (simp add: less_degree_pdevs_of_list_imp_less_length pdevs_apply_pdevs_of_list) also have "\ = (\iR (sum_list xs))" by (simp add: r less_degree_pdevs_of_list_imp_less_length) also have "\ = (\iR (sum_list xs)" by (simp add: algebra_simps scaleR_sum_left) finally have eq: "pdevs_val f ?coll = (\iR (sum_list xs)" (is "_ = ?e *\<^sub>R _") . have "abs (pdevs_val f ?coll) < tdev ?coll" using assms(4) by (intro abs_pdevs_val_less_tdev) (auto simp: Pi_iff less_imp_le deg_pos) also have "\ = sum_list (map abs xs)" using assms by simp also note eq finally have less: "\?e\ *\<^sub>R abs (sum_list xs) < sum_list (map abs xs)" by (simp add: abs_scaleR) also have "\sum_list xs\ = sum_list (map abs xs)" using coll \x \ 0\ nlex by (rule abs_sum_list_coll) finally have "?e \ {-1 <..< 1}" by (auto simp add: less_le scaleR_le_self_cancel) thus ?thesis using eq .. qed subsection \Independent Generators\ fun independent_pdevs::"point list \ point list" where "independent_pdevs [] = []" | "independent_pdevs (x#xs) = (let (cs, is) = List.partition (coll 0 x) xs; s = x + sum_list cs in (if s = 0 then [] else [s]) @ independent_pdevs is)" lemma in_set_independent_pdevsE: assumes "y \ set (independent_pdevs xs)" obtains x where "x\set xs" "coll 0 x y" proof atomize_elim show "\x. x \ set xs \ coll 0 x y" using assms proof (induct xs rule: independent_pdevs.induct) case 1 thus ?case by simp next case (2 z zs) let ?c1 = "y = z + sum_list (filter (coll 0 z) zs)" let ?c2 = "y \ set (independent_pdevs (filter (Not \ coll 0 z) zs))" from 2 have "?c1 \ ?c2" by (auto simp: Let_def split: if_split_asm) thus ?case proof assume ?c2 hence "y \ set (independent_pdevs (snd (partition (coll 0 z) zs)))" by simp from 2(1)[OF refl prod.collapse refl this] show ?case by auto next assume y: ?c1 show ?case unfolding y by (rule exI[where x="z"]) (auto intro!: coll_add coll_sum_list ) qed qed qed lemma in_set_independent_pdevs_nonzero: "x \ set (independent_pdevs xs) \ x \ 0" proof (induct xs rule: independent_pdevs.induct) case (2 y ys) from 2(1)[OF refl prod.collapse refl] 2(2) show ?case by (auto simp: Let_def split: if_split_asm) qed simp lemma independent_pdevs_pairwise_non_coll: assumes "x \ set (independent_pdevs xs)" assumes "y \ set (independent_pdevs xs)" assumes "\x. x \ set xs \ x \ 0" assumes "x \ y" shows "\ coll 0 x y" using assms proof (induct xs rule: independent_pdevs.induct) case 1 thus ?case by simp next case (2 z zs) from 2 have "z \ 0" by simp from 2(2) have "x \ 0" by (rule in_set_independent_pdevs_nonzero) from 2(3) have "y \ 0" by (rule in_set_independent_pdevs_nonzero) let ?c = "filter (coll 0 z) zs" let ?nc = "filter (Not \ coll 0 z) zs" { assume "x \ set (independent_pdevs ?nc)" "y \ set (independent_pdevs ?nc)" hence "\coll 0 x y" by (intro 2(1)[OF refl prod.collapse refl _ _ 2(4) 2(5)]) auto } note IH = this { fix x assume "x \ 0" "z + sum_list ?c \ 0" "coll 0 x (z + sum_list ?c)" hence "x \ set (independent_pdevs ?nc)" using sum_list_filter_coll_ex_scale[OF \z \ 0\, of "z#zs"] by (auto elim!: in_set_independent_pdevsE simp: coll_commute) (metis (no_types) \x \ 0\ coll_scale coll_scaleR) } note nc = this from 2(2,3,4,5) nc[OF \x \ 0\] nc[OF \y \ 0\] show ?case by (auto simp: Let_def IH coll_commute split: if_split_asm) qed lemma distinct_independent_pdevs[simp]: shows "distinct (independent_pdevs xs)" proof (induct xs rule: independent_pdevs.induct) case 1 thus ?case by simp next case (2 x xs) let ?is = "independent_pdevs (filter (Not \ coll 0 x) xs)" have "distinct ?is" by (rule 2) (auto intro!: 2) thus ?case proof (clarsimp simp add: Let_def) let ?s = "x + sum_list (filter (coll 0 x) xs)" assume s: "?s \ 0" "?s \ set ?is" from in_set_independent_pdevsE[OF s(2)] obtain y where y: "y \ set (filter (Not \ coll 0 x) xs)" "coll 0 y (x + sum_list (filter (coll 0 x) xs))" by auto { assume "y = 0 \ x = 0 \ sum_list (filter (coll 0 x) xs) = 0" hence False using s y by (auto simp: coll_commute) } moreover { assume "y \ 0" "x \ 0" "sum_list (filter (coll 0 x) xs) \ 0" "sum_list (filter (coll 0 x) xs) + x \ 0" have *: "coll 0 (sum_list (filter (coll 0 x) xs)) x" by (auto intro!: coll_sum_list simp: coll_commute) have "coll 0 y (sum_list (filter (coll 0 x) xs) + x)" using s y by (simp add: add.commute) hence "coll 0 y x" using * by (rule coll_add_trans) fact+ hence False using s y by (simp add: coll_commute) } ultimately show False using s y by (auto simp: add.commute) qed qed lemma in_set_independent_pdevs_invariant_nlex: "x \ set (independent_pdevs xs) \ (\x. x \ set xs \ lex x 0) \ (\x. x \ set xs \ x \ 0) \ Counterclockwise_2D_Arbitrary.lex x 0" proof (induction xs arbitrary: x rule: independent_pdevs.induct ) case 1 thus ?case by simp next case (2 y ys) from 2 have "y \ 0" by auto from 2(2) have "x \ set (independent_pdevs (filter (Not \ coll 0 y) ys)) \ x = y + sum_list (filter (coll 0 y) ys)" by (auto simp: Let_def split: if_split_asm) thus ?case proof assume "x \ set (independent_pdevs (filter (Not \ coll 0 y) ys))" from 2(1)[OF refl prod.collapse refl, simplified, OF this 2(3,4)] show ?case by simp next assume "x = y + sum_list (filter (coll 0 y) ys)" also have "lex \ 0" by (force intro: nlex_add nlex_sum simp: sum_list_sum_nth dest: nth_mem intro: 2(3)) finally show ?case . qed qed lemma pdevs_val_independent_pdevs2: assumes "e \ UNIV \ I" shows "\e'. e' \ UNIV \ I \ pdevs_val e (pdevs_of_list (independent_pdevs xs)) = pdevs_val e' (pdevs_of_list xs)" using assms proof (induct xs arbitrary: e rule: independent_pdevs.induct) case 1 thus ?case by force next case (2 x xs) let ?coll = "(filter (coll 0 x) (x#xs))" let ?ncoll = "(filter (Not o coll 0 x) (x#xs))" let ?e0 = "if sum_list ?coll = 0 then e else e \ (+) (Suc 0)" have "pdevs_val e (pdevs_of_list (independent_pdevs (x#xs))) = e 0 *\<^sub>R (sum_list ?coll) + pdevs_val ?e0 (pdevs_of_list (independent_pdevs ?ncoll))" (is "_ = ?vc + ?vnc") by (simp add: Let_def) also have e_split: "(\_. e 0) \ UNIV \ I" "?e0 \ UNIV \ I" using 2(2) by auto from 2(1)[OF refl prod.collapse refl e_split(2)] obtain e' where e': "e' \ UNIV \ I" and "?vnc = pdevs_val e' (pdevs_of_list ?ncoll)" by (auto simp add: o_def) note this(2) also have "?vc = pdevs_val (\_. e 0) (pdevs_of_list ?coll)" by (simp add: pdevs_val_const_pdevs_of_list) also from pdevs_val_pdevs_of_list_append[OF e_split(1) e'] obtain e'' where e'': "e'' \ UNIV \ I" and "pdevs_val (\_. e 0) (pdevs_of_list ?coll) + pdevs_val e' (pdevs_of_list ?ncoll) = pdevs_val e'' (pdevs_of_list (?coll @ ?ncoll))" by metis note this(2) also from pdevs_val_perm[OF partition_permI e'', of "coll 0 x" "x#xs"] obtain e''' where e''': "e''' \ UNIV \ I" and "\ = pdevs_val e''' (pdevs_of_list (x # xs))" by metis note this(2) finally show ?case using e''' by auto qed lemma list_all_filter: "list_all p (filter p xs)" by (induct xs) auto lemma pdevs_val_independent_pdevs: assumes "list_all (\x. lex x 0) xs" assumes "list_all (\x. x \ 0) xs" assumes "e \ UNIV \ {-1 .. 1}" shows "\e'. e' \ UNIV \ {-1 .. 1} \ pdevs_val e (pdevs_of_list xs) = pdevs_val e' (pdevs_of_list (independent_pdevs xs))" using assms(1,2,3) proof (induct xs arbitrary: e rule: independent_pdevs.induct) case 1 thus ?case by force next case (2 x xs) let ?coll = "(filter (coll 0 x) (x#xs))" let ?ncoll = "(filter (Not o coll 0 x) xs)" from 2 have "x \ 0" by simp from pdevs_val_partition[OF 2(4), of "x#xs" "coll 0 x"] obtain f g where part: "pdevs_val e (pdevs_of_list (x # xs)) = pdevs_val f (pdevs_of_list ?coll) + pdevs_val g (pdevs_of_list (filter (Not o coll 0 x) (x#xs)))" and f: "f \ UNIV \ {-1 .. 1}" and g: "g \ UNIV \ {-1 .. 1}" by blast note part also have "list_all (\x. lex x 0) (filter (coll 0 x) (x#xs))" using 2(2) by (auto simp: inner_prod_def list_all_iff) from pdevs_val_coll[OF list_all_filter this \x \ 0\ f] obtain f' where "pdevs_val f (pdevs_of_list ?coll) = f' *\<^sub>R sum_list (filter (coll 0 x) (x#xs))" and f': "f' \ {-1 .. 1}" by blast note this(1) also have "filter (Not o coll 0 x) (x#xs) = ?ncoll" by simp also from 2(2) have "list_all (\x. lex x 0) ?ncoll" "list_all (\x. x \ 0) ?ncoll" by (auto simp: list_all_iff) from 2(1)[OF refl partition_filter_conv[symmetric] refl this g] obtain g' where "pdevs_val g (pdevs_of_list ?ncoll) = pdevs_val g' (pdevs_of_list (independent_pdevs ?ncoll))" and g': "g' \ UNIV \ {-1 .. 1}" by metis note this(1) also define h where "h = (if sum_list ?coll \ 0 then rec_nat f' (\i _. g' i) else g')" have "f' *\<^sub>R sum_list ?coll + pdevs_val g' (pdevs_of_list (independent_pdevs ?ncoll)) = pdevs_val h (pdevs_of_list (independent_pdevs (x#xs)))" by (simp add: h_def o_def Let_def) finally have "pdevs_val e (pdevs_of_list (x # xs)) = pdevs_val h (pdevs_of_list (independent_pdevs (x # xs)))" . moreover have "h \ UNIV \ {-1 .. 1}" proof fix i show "h i \ {-1 .. 1}" using f' g' by (cases i) (auto simp: h_def) qed ultimately show ?case by blast qed lemma pdevs_val_independent_pdevs_strict: assumes "list_all (\x. lex x 0) xs" assumes "list_all (\x. x \ 0) xs" assumes "e \ UNIV \ {-1 <..< 1}" shows "\e'. e' \ UNIV \ {-1 <..< 1} \ pdevs_val e (pdevs_of_list xs) = pdevs_val e' (pdevs_of_list (independent_pdevs xs))" using assms(1,2,3) proof (induct xs arbitrary: e rule: independent_pdevs.induct) case 1 thus ?case by force next case (2 x xs) let ?coll = "(filter (coll 0 x) (x#xs))" let ?ncoll = "(filter (Not o coll 0 x) xs)" from 2 have "x \ 0" by simp from pdevs_val_partition[OF 2(4), of "x#xs" "coll 0 x"] obtain f g where part: "pdevs_val e (pdevs_of_list (x # xs)) = pdevs_val f (pdevs_of_list ?coll) + pdevs_val g (pdevs_of_list (filter (Not o coll 0 x) (x#xs)))" and f: "f \ UNIV \ {-1 <..< 1}" and g: "g \ UNIV \ {-1 <..< 1}" by blast note part also have "list_all (\x. lex x 0) (filter (coll 0 x) (x#xs))" using 2(2) by (auto simp: inner_prod_def list_all_iff) from pdevs_val_coll_strict[OF list_all_filter this \x \ 0\ f] obtain f' where "pdevs_val f (pdevs_of_list ?coll) = f' *\<^sub>R sum_list (filter (coll 0 x) (x#xs))" and f': "f' \ {-1 <..< 1}" by blast note this(1) also have "filter (Not o coll 0 x) (x#xs) = ?ncoll" by simp also from 2(2) have "list_all (\x. lex x 0) ?ncoll" "list_all (\x. x \ 0) ?ncoll" by (auto simp: list_all_iff) from 2(1)[OF refl partition_filter_conv[symmetric] refl this g] obtain g' where "pdevs_val g (pdevs_of_list ?ncoll) = pdevs_val g' (pdevs_of_list (independent_pdevs ?ncoll))" and g': "g' \ UNIV \ {-1 <..< 1}" by metis note this(1) also define h where "h = (if sum_list ?coll \ 0 then rec_nat f' (\i _. g' i) else g')" have "f' *\<^sub>R sum_list ?coll + pdevs_val g' (pdevs_of_list (independent_pdevs ?ncoll)) = pdevs_val h (pdevs_of_list (independent_pdevs (x#xs)))" by (simp add: h_def o_def Let_def) finally have "pdevs_val e (pdevs_of_list (x # xs)) = pdevs_val h (pdevs_of_list (independent_pdevs (x # xs)))" . moreover have "h \ UNIV \ {-1 <..< 1}" proof fix i show "h i \ {-1 <..< 1}" using f' g' by (cases i) (auto simp: h_def) qed ultimately show ?case by blast qed lemma sum_list_independent_pdevs: "sum_list (independent_pdevs xs) = sum_list xs" proof (induct xs rule: independent_pdevs.induct) case (2 y ys) from 2[OF refl prod.collapse refl] show ?case using sum_list_partition[of "coll 0 y" ys, symmetric] by (auto simp: Let_def) qed simp lemma independent_pdevs_eq_Nil_iff: "list_all (\x. lex x 0) xs \ list_all (\x. x \ 0) xs \ independent_pdevs xs = [] \ xs = []" proof (induct xs) case Nil thus ?case by simp next case (Cons x xs) from Cons(2) have "list_all (\x. lex x 0) (x # filter (coll 0 x) xs)" by (auto simp: list_all_iff) from sum_list_nlex_eq_zero_iff[OF this] Cons(3) show ?case by (auto simp: list_all_iff) qed subsection \Independent Oriented Generators\ definition "inl p = independent_pdevs (map snd (list_of_pdevs (nlex_pdevs p)))" lemma distinct_inl[simp]: "distinct (inl (snd X))" by (auto simp: inl_def) lemma in_set_inl_nonzero: "x \ set (inl xs) \ x \ 0" by (auto simp: inl_def in_set_independent_pdevs_nonzero) lemma inl_ncoll: "y \ set (inl (snd X)) \ z \ set (inl (snd X)) \ y \ z \ \coll 0 y z" unfolding inl_def by (rule independent_pdevs_pairwise_non_coll, assumption+) (auto simp: inl_def list_of_pdevs_nonzero) lemma in_set_inl_lex: "x \ set (inl xs) \ lex x 0" by (auto simp: inl_def list_of_pdevs_def dest!: in_set_independent_pdevs_invariant_nlex split: if_split_asm) interpretation ccw0: linorder_list "ccw 0" "set (inl (snd X))" proof unfold_locales fix a b c show "a \ b \ ccw 0 a b \ ccw 0 b a" by (metis UNIV_I ccw_self(1) nondegenerate) assume a: "a \ set (inl (snd X))" show "ccw 0 a a" by simp assume b: "b \ set (inl (snd X))" show "ccw 0 a b \ ccw 0 b a \ a = b" by (metis ccw_self(1) in_set_inl_nonzero mem_Collect_eq not_ccw_eq a b) assume c: "c \ set (inl (snd X))" assume distinct: "a \ b" "b \ c" "a \ c" assume ab: "ccw 0 a b" and bc: "ccw 0 b c" show "ccw 0 a c" using a b c ab bc proof (cases "a = (0, 1)" "b = (0, 1)" "c = (0, 1)" rule: case_split[case_product case_split[case_product case_split]]) assume nu: "a \ (0, 1)" "b \ (0, 1)" "c \ (0, 1)" have "distinct5 a b c (0, 1) 0" "in5 UNIV a b c (0, 1) 0" using a b c distinct nu by (simp_all add: in_set_inl_nonzero) moreover have "ccw 0 (0, 1) a" "ccw 0 (0, 1) b" "ccw 0 (0, 1) c" by (auto intro!: nlex_ccw_left in_set_inl_lex a b c) ultimately show ?thesis using ab bc by (rule ccw.transitive[where S=UNIV and s="(0, 1)"]) next assume "a \ (0, 1)" "b = (0, 1)" "c \ (0, 1)" thus ?thesis using ccw_switch23 in_set_inl_lex inl_ncoll nlex_ccw_left a b ab by blast next assume "a \ (0, 1)" "b \ (0, 1)" "c = (0, 1)" thus ?thesis using ccw_switch23 in_set_inl_lex inl_ncoll nlex_ccw_left b c bc by blast qed (auto simp add: nlex_ccw_left in_set_inl_lex) qed lemma sorted_inl: "ccw.sortedP 0 (ccw.selsort 0 (inl (snd X)))" by (rule ccw0.sortedP_selsort) auto lemma sorted_scaled_inl: "ccw.sortedP 0 (map ((*\<^sub>R) 2) (ccw.selsort 0 (inl (snd X))))" using sorted_inl by (rule ccw_sorted_scaleR) simp lemma distinct_selsort_inl: "distinct (ccw.selsort 0 (inl (snd X)))" by simp lemma distinct_map_scaleRI: fixes xs::"'a::real_vector list" shows "distinct xs \ c \ 0 \ distinct (map ((*\<^sub>R) c) xs)" by (induct xs) auto lemma distinct_scaled_inl: "distinct (map ((*\<^sub>R) 2) (ccw.selsort 0 (inl (snd X))))" using distinct_selsort_inl by (rule distinct_map_scaleRI) simp lemma ccw'_sortedP_scaled_inl: "ccw'.sortedP 0 (map ((*\<^sub>R) 2) (ccw.selsort 0 (inl (snd X))))" using ccw_sorted_implies_ccw'_sortedP by (rule ccw'_sorted_scaleR) (auto simp: sorted_inl inl_ncoll) lemma pdevs_val_pdevs_of_list_inl2E: assumes "e \ UNIV \ {-1 .. 1}" obtains e' where "pdevs_val e X = pdevs_val e' (pdevs_of_list (inl X))" "e' \ UNIV \ {-1 .. 1}" proof - let ?l = "map snd (list_of_pdevs (nlex_pdevs X))" have l: "list_all (\x. Counterclockwise_2D_Arbitrary.lex x 0) ?l" "list_all (\x. x \ 0) (map snd (list_of_pdevs (nlex_pdevs X)))" by (auto simp: list_all_iff list_of_pdevs_def) from pdevs_val_nlex_pdevs[OF assms(1)] obtain e' where "e' \ UNIV \ {-1 .. 1}" "pdevs_val e X = pdevs_val e' (nlex_pdevs X)" by auto note this(2) also from pdevs_val_of_list_of_pdevs2[OF \e' \ _\] obtain e'' where "e'' \ UNIV \ {-1 .. 1}" "\ = pdevs_val e'' (pdevs_of_list ?l)" by metis note this(2) also from pdevs_val_independent_pdevs[OF l \e'' \ _\] obtain e''' where "e''' \ UNIV \ {-1 .. 1}" and "\ = pdevs_val e''' (pdevs_of_list (independent_pdevs ?l))" by metis note this(2) also have "\ = pdevs_val e''' (pdevs_of_list (inl X))" by (simp add: inl_def) finally have "pdevs_val e X = pdevs_val e''' (pdevs_of_list (inl X))" . thus thesis using \e''' \ _\ .. qed lemma pdevs_val_pdevs_of_list_inlE: assumes "e \ UNIV \ I" "uminus ` I = I" "0 \ I" obtains e' where "pdevs_val e (pdevs_of_list (inl X)) = pdevs_val e' X" "e' \ UNIV \ I" proof - let ?l = "map snd (list_of_pdevs (nlex_pdevs X))" have "pdevs_val e (pdevs_of_list (inl X)) = pdevs_val e (pdevs_of_list (independent_pdevs ?l))" by (simp add: inl_def) also from pdevs_val_independent_pdevs2[OF \e \ _\] obtain e' where "pdevs_val e (pdevs_of_list (independent_pdevs ?l)) = pdevs_val e' (pdevs_of_list ?l)" and "e' \ UNIV \ I" by auto note this(1) also from pdevs_val_of_list_of_pdevs[OF \e' \ _\ \0 \ I\, of "nlex_pdevs X"] obtain e'' where "pdevs_val e' (pdevs_of_list ?l) = pdevs_val e'' (nlex_pdevs X)" and "e'' \ UNIV \ I" by metis note this(1) also from pdevs_val_nlex_pdevs2[OF \e'' \ _\ \_ = I\] obtain e''' where "pdevs_val e'' (nlex_pdevs X) = pdevs_val e''' X" "e''' \ UNIV \ I" by metis note this(1) finally have "pdevs_val e (pdevs_of_list (inl X)) = pdevs_val e''' X" . thus ?thesis using \e''' \ UNIV \ I\ .. qed lemma sum_list_nlex_eq_sum_list_inl: "sum_list (map snd (list_of_pdevs (nlex_pdevs X))) = sum_list (inl X)" by (auto simp: inl_def sum_list_list_of_pdevs sum_list_independent_pdevs) lemma Affine_inl: "Affine (fst X, pdevs_of_list (inl (snd X))) = Affine X" by (auto simp: Affine_def valuate_def aform_val_def elim: pdevs_val_pdevs_of_list_inlE[of _ _ "snd X"] pdevs_val_pdevs_of_list_inl2E[of _ "snd X"]) subsection \Half Segments\ definition half_segments_of_aform::"point aform \ (point*point) list" where "half_segments_of_aform X = (let x0 = lowest_vertex (fst X, nlex_pdevs (snd X)) in polychain_of x0 (map ((*\<^sub>R) 2) (ccw.selsort 0 (inl (snd X)))))" lemma subsequent_half_segments: fixes X assumes "Suc i < length (half_segments_of_aform X)" shows "snd (half_segments_of_aform X ! i) = fst (half_segments_of_aform X ! Suc i)" using assms by (cases i) (auto simp: half_segments_of_aform_def Let_def polychain_of_subsequent_eq) lemma polychain_half_segments_of_aform: "polychain (half_segments_of_aform X)" by (auto simp: subsequent_half_segments intro!: polychainI) lemma fst_half_segments: "half_segments_of_aform X \ [] \ fst (half_segments_of_aform X ! 0) = lowest_vertex (fst X, nlex_pdevs (snd X))" by (auto simp: half_segments_of_aform_def Let_def o_def split_beta') lemma nlex_half_segments_of_aform: "(a, b) \ set (half_segments_of_aform X) \ lex b a" by (auto simp: half_segments_of_aform_def prod_eq_iff lex_def dest!: in_set_polychain_ofD in_set_inl_lex) lemma ccw_half_segments_of_aform_all: assumes cd: "(c, d) \ set (half_segments_of_aform X)" shows "list_all (\(xi, xj). ccw xi xj c \ ccw xi xj d) (half_segments_of_aform X)" proof - have "list_all (\(xi, xj). ccw xi xj (fst (c, d)) \ ccw xi xj (snd (c, d))) (polychain_of (lowest_vertex (fst X, nlex_pdevs (snd X))) ((map ((*\<^sub>R) 2) (linorder_list0.selsort (ccw 0) (inl (snd X))))))" using ccw'_sortedP_scaled_inl cd[unfolded half_segments_of_aform_def Let_def] by (rule polychain_of_ccw_conjunction) thus ?thesis unfolding half_segments_of_aform_def[unfolded Let_def, symmetric] fst_conv snd_conv . qed lemma ccw_half_segments_of_aform: assumes ij: "(xi, xj) \ set (half_segments_of_aform X)" assumes c: "(c, d) \ set (half_segments_of_aform X)" shows "ccw xi xj c" "ccw xi xj d" using ccw_half_segments_of_aform_all[OF c] ij by (auto simp add: list_all_iff) lemma half_segments_of_aform1: assumes ch: "x \ convex hull set (map fst (half_segments_of_aform X))" assumes ab: "(a, b) \ set (half_segments_of_aform X)" shows "ccw a b x" using finite_set _ ch proof (rule ccw.convex_hull) fix c assume "c \ set (map fst (half_segments_of_aform X))" then obtain d where "(c, d) \ set (half_segments_of_aform X)" by auto with ab show "ccw a b c" by (rule ccw_half_segments_of_aform(1)) qed (insert ab, simp add: nlex_half_segments_of_aform) lemma half_segments_of_aform2: assumes ch: "x \ convex hull set (map snd (half_segments_of_aform X))" assumes ab: "(a, b) \ set (half_segments_of_aform X)" shows "ccw a b x" using finite_set _ ch proof (rule ccw.convex_hull) fix d assume "d \ set (map snd (half_segments_of_aform X))" then obtain c where "(c, d) \ set (half_segments_of_aform X)" by auto with ab show "ccw a b d" by (rule ccw_half_segments_of_aform(2)) qed (insert ab, simp add: nlex_half_segments_of_aform) lemma in_set_half_segments_of_aform_aform_valE: assumes "(x2, y2) \ set (half_segments_of_aform X)" obtains e where "y2 = aform_val e X" "e \ UNIV \ {-1 .. 1}" proof - from assms obtain d where "y2 = lowest_vertex (fst X, nlex_pdevs (snd X)) + sum_list (take (Suc d) (map ((*\<^sub>R) 2) (ccw.selsort 0 (inl (snd X)))))" by (auto simp: half_segments_of_aform_def elim!: in_set_polychain_of_imp_sum_list) also have "lowest_vertex (fst X, nlex_pdevs (snd X)) = fst X - sum_list (map snd (list_of_pdevs (nlex_pdevs (snd X))))" by (simp add: lowest_vertex_def) also have "sum_list (map snd (list_of_pdevs (nlex_pdevs (snd X)))) = pdevs_val (\_. 1) (nlex_pdevs (snd X))" by (auto simp: pdevs_val_sum_list) also have "sum_list (take (Suc d) (map ((*\<^sub>R) 2) (ccw.selsort 0 (inl (snd X))))) = pdevs_val (\i. if i \ d then 2 else 0) (pdevs_of_list (ccw.selsort 0 (inl (snd X))))" (is "_ = pdevs_val ?e _") by (subst sum_list_take_pdevs_val_eq) (auto simp: pdevs_val_sum if_distrib pdevs_apply_pdevs_of_list degree_pdevs_of_list_scaleR intro!: sum.cong ) also obtain e'' where "\ = pdevs_val e'' (pdevs_of_list (inl (snd X)))" "e'' \ UNIV \ {0..2}" by (auto intro: pdevs_val_selsort_ccw2[of "inl (snd X)" ?e "{0 .. 2}"]) note this(1) also note inl_def also let ?l = "map snd (list_of_pdevs (nlex_pdevs (snd X)))" from pdevs_val_independent_pdevs2[OF \e'' \ _\] obtain e''' where "pdevs_val e'' (pdevs_of_list (independent_pdevs ?l)) = pdevs_val e''' (pdevs_of_list ?l)" and "e''' \ UNIV \ {0..2}" by auto note this(1) also have "0 \ {0 .. 2::real}" by simp from pdevs_val_of_list_of_pdevs[OF \e''' \ _\ this, of "nlex_pdevs (snd X)"] obtain e'''' where "pdevs_val e''' (pdevs_of_list ?l) = pdevs_val e'''' (nlex_pdevs (snd X))" and "e'''' \ UNIV \ {0 .. 2}" by metis note this(1) finally have "y2 = fst X + (pdevs_val e'''' (nlex_pdevs (snd X)) - pdevs_val (\_. 1) (nlex_pdevs (snd X)))" by simp also have "pdevs_val e'''' (nlex_pdevs (snd X)) - pdevs_val (\_. 1) (nlex_pdevs (snd X)) = pdevs_val (\i. e'''' i - 1) (nlex_pdevs (snd X))" by (simp add: pdevs_val_minus) also have "(\i. e'''' i - 1) \ UNIV \ {-1 .. 1}" using \e'''' \ _\ by auto from pdevs_val_nlex_pdevs2[OF this] obtain f where "f \ UNIV \ {-1 .. 1}" and "pdevs_val (\i. e'''' i - 1) (nlex_pdevs (snd X)) = pdevs_val f (snd X)" by auto note this(2) finally have "y2 = aform_val f X" by (simp add: aform_val_def) thus ?thesis using \f \ _\ .. qed lemma fst_hd_half_segments_of_aform: assumes "half_segments_of_aform X \ []" shows "fst (hd (half_segments_of_aform X)) = lowest_vertex (fst X, (nlex_pdevs (snd X)))" using assms by (auto simp: half_segments_of_aform_def Let_def fst_hd_polychain_of) lemma "linorder_list0.sortedP (ccw' (lowest_vertex (fst X, nlex_pdevs (snd X)))) (map snd (half_segments_of_aform X))" (is "linorder_list0.sortedP (ccw' ?x0) ?ms") unfolding half_segments_of_aform_def Let_def by (rule ccw'_sortedP_polychain_of_snd) (rule ccw'_sortedP_scaled_inl) lemma rev_zip: "length xs = length ys \ rev (zip xs ys) = zip (rev xs) (rev ys)" by (induct xs ys rule: list_induct2) auto lemma zip_upt_self_aux: "zip [0..i. (i, xs ! i)) [0.. UNIV \ {-1 <..< 1}" assumes "seg \ set (half_segments_of_aform X)" assumes "length (half_segments_of_aform X) \ 1" shows "ccw' (fst seg) (snd seg) (aform_val e X)" using assms unfolding half_segments_of_aform_def Let_def proof - have len: "length (map ((*\<^sub>R) 2) (linorder_list0.selsort (ccw 0) (inl (snd X)))) \ 1" using assms by (auto simp: half_segments_of_aform_def) have "aform_val e X = fst X + pdevs_val e (snd X)" by (simp add: aform_val_def) also obtain e' where "e' \ UNIV \ {-1 <..< 1}" "pdevs_val e (snd X) = pdevs_val e' (nlex_pdevs (snd X))" using pdevs_val_nlex_pdevs[OF \e \ _\] by auto note this(2) also obtain e'' where "e'' \ UNIV \ {-1 <..< 1}" "\ = pdevs_val e'' (pdevs_of_list (map snd (list_of_pdevs (nlex_pdevs (snd X)))))" by (metis pdevs_val_of_list_of_pdevs2[OF \e' \ _\]) note this(2) also obtain e''' where "e''' \ UNIV \ {-1 <..< 1}" "\ = pdevs_val e''' (pdevs_of_list (inl (snd X)))" unfolding inl_def using pdevs_val_independent_pdevs_strict[OF list_all_list_of_pdevsI, OF lex_nlex_pdevs list_of_pdevs_all_nonzero \e'' \ _\] by metis note this(2) also from pdevs_val_selsort_ccw[OF distinct_inl \e''' \ _\] obtain f where "f \ UNIV \ {-1 <..< 1}" "\ = pdevs_val f (pdevs_of_list (linorder_list0.selsort (ccw 0) (inl (snd X))))" (is "_ = pdevs_val _ (pdevs_of_list ?sl)") by metis note this(2) also have "\ = pdevs_val (\i. f i + 1) (pdevs_of_list ?sl) + lowest_vertex (fst X, nlex_pdevs (snd X)) - fst X" proof - have "sum_list (dense_list_of_pdevs (nlex_pdevs (snd X))) = sum_list (dense_list_of_pdevs (pdevs_of_list (ccw.selsort 0 (inl (snd X)))))" by (subst dense_list_of_pdevs_pdevs_of_list) (auto simp: in_set_independent_pdevs_nonzero dense_list_of_pdevs_pdevs_of_list inl_def sum_list_distinct_selsort sum_list_independent_pdevs sum_list_list_of_pdevs) thus ?thesis by (auto simp add: pdevs_val_add lowest_vertex_def algebra_simps pdevs_val_sum_list sum_list_list_of_pdevs in_set_inl_nonzero dense_list_of_pdevs_pdevs_of_list) qed also have "pdevs_val (\i. f i + 1) (pdevs_of_list ?sl) = pdevs_val (\i. 1/2 * (f i + 1)) (pdevs_of_list (map ((*\<^sub>R) 2) ?sl))" (is "_ = pdevs_val ?f' (pdevs_of_list ?ssl)") by (subst pdevs_val_cmul) (simp add: pdevs_of_list_map_scaleR) also have "distinct ?ssl" "?f' \ UNIV \ {0<..<1}" using \f \ _\ by (auto simp: distinct_map_scaleRI Pi_iff algebra_simps real_0_less_add_iff) from pdevs_of_list_sum[OF this] obtain g where "g \ UNIV \ {0<..<1}" "pdevs_val ?f' (pdevs_of_list ?ssl) = (\P\set ?ssl. g P *\<^sub>R P)" by blast note this(2) finally have "aform_val e X = lowest_vertex (fst X, nlex_pdevs (snd X)) + (\P\set ?ssl. g P *\<^sub>R P)" by simp also have "ccw' (fst seg) (snd seg) \" using \g \ _\ _ len \seg \ _\[unfolded half_segments_of_aform_def Let_def] by (rule in_polychain_of_ccw) (simp add: ccw'_sortedP_scaled_inl) finally show ?thesis . qed lemma half_segments_of_aform_strict_all: assumes "e \ UNIV \ {-1 <..< 1}" assumes "length (half_segments_of_aform X) \ 1" shows "list_all (\seg. ccw' (fst seg) (snd seg) (aform_val e X)) (half_segments_of_aform X)" using assms by (auto intro!: half_segments_of_aform_strict simp: list_all_iff) lemma list_allD: "list_all P xs \ x \ set xs \ P x" by (auto simp: list_all_iff) lemma minus_one_less_multI: fixes e x::real shows "- 1 \ e \ 0 < x \ x < 1 \ - 1 < e * x" by (metis abs_add_one_gt_zero abs_real_def le_less_trans less_not_sym mult_less_0_iff mult_less_cancel_left1 real_0_less_add_iff) lemma less_one_multI: fixes e x::real shows "e \ 1 \ 0 < x \ x < 1 \ e * x < 1" by (metis (erased, hide_lams) less_eq_real_def monoid_mult_class.mult.left_neutral mult_strict_mono zero_less_one) lemma ccw_half_segments_of_aform_strictI: assumes "e \ UNIV \ {-1 <..< 1}" assumes "(s1, s2) \ set (half_segments_of_aform X)" assumes "length (half_segments_of_aform X) \ 1" assumes "x = (aform_val e X)" shows "ccw' s1 s2 x" using half_segments_of_aform_strict[OF assms(1-3)] assms(4) by simp lemma ccw'_sortedP_subsequent: assumes "Suc i < length xs" "ccw'.sortedP 0 (map dirvec xs)" "fst (xs ! Suc i) = snd (xs ! i)" shows "ccw' (fst (xs ! i)) (snd (xs ! i)) (snd (xs ! Suc i))" using assms proof (induct xs arbitrary: i) case Nil thus ?case by simp next case (Cons x xs) thus ?case by (auto simp: nth_Cons dirvec_minus split: nat.split elim!: ccw'.sortedP_Cons) (metis (no_types, lifting) ccw'.renormalize length_greater_0_conv nth_mem prod.case_eq_if) qed lemma ccw'_sortedP_uminus: "ccw'.sortedP 0 xs \ ccw'.sortedP 0 (map uminus xs)" by (induct xs) (auto intro!: ccw'.sortedP.Cons elim!: ccw'.sortedP_Cons simp del: uminus_Pair) lemma subsequent_half_segments_ccw: fixes X assumes "Suc i < length (half_segments_of_aform X)" shows "ccw' (fst (half_segments_of_aform X ! i)) (snd (half_segments_of_aform X ! i)) (snd (half_segments_of_aform X ! Suc i))" using assms by (intro ccw'_sortedP_subsequent ) (auto simp: subsequent_half_segments half_segments_of_aform_def sorted_inl polychain_of_subsequent_eq intro!: ccw_sorted_implies_ccw'_sortedP[OF inl_ncoll] ccw'_sorted_scaleR) lemma convex_polychain_half_segments_of_aform: "convex_polychain (half_segments_of_aform X)" proof cases assume "length (half_segments_of_aform X) = 1" thus ?thesis by (auto simp: length_Suc_conv convex_polychain_def polychain_def) next assume len: "length (half_segments_of_aform X) \ 1" show ?thesis by (rule convex_polychainI) (simp_all add: polychain_half_segments_of_aform subsequent_half_segments_ccw ccw'_def[symmetric]) qed lemma hd_distinct_neq_last: "distinct xs \ length xs > 1 \ hd xs \ last xs" by (metis One_nat_def add_Suc_right distinct.simps(2) last.simps last_in_set less_irrefl list.exhaust list.sel(1) list.size(3) list.size(4) add.right_neutral nat_neq_iff not_less0) lemma ccw_hd_last_half_segments_dirvec: assumes "length (half_segments_of_aform X) > 1" shows "ccw' 0 (dirvec (hd (half_segments_of_aform X))) (dirvec (last (half_segments_of_aform X)))" proof - let ?i = "ccw.selsort 0 (inl (snd X))" let ?s = "map ((*\<^sub>R) 2) (ccw.selsort 0 (inl (snd X)))" from assms have l: "1 < length (inl (snd X))" "inl (snd X) \ []" using assms by (auto simp add: half_segments_of_aform_def) hence "hd ?i \ set ?i" "last ?i \ set ?i" by (auto intro!: hd_in_set last_in_set simp del: ccw.set_selsort) hence "\coll 0 (hd ?i) (last ?i)" using l by (intro inl_ncoll[of _ X]) (auto simp: hd_distinct_neq_last) hence "\coll 0 (hd ?s) (last ?s)" using l by (auto simp: hd_map last_map) hence "ccw' 0 (hd (map ((*\<^sub>R) 2) (linorder_list0.selsort (ccw 0) (inl (snd X))))) (last (map ((*\<^sub>R) 2) (linorder_list0.selsort (ccw 0) (inl (snd X)))))" using assms by (auto simp add: half_segments_of_aform_def intro!: sorted_inl ccw_sorted_scaleR ccw.hd_last_sorted ccw_ncoll_imp_ccw) with assms show ?thesis by (auto simp add: half_segments_of_aform_def Let_def dirvec_hd_polychain_of dirvec_last_polychain_of length_greater_0_conv[symmetric] simp del: polychain_of.simps length_greater_0_conv split: if_split_asm) qed lemma map_fst_half_segments_aux_eq: "[] \ half_segments_of_aform X \ map fst (half_segments_of_aform X) = fst (hd (half_segments_of_aform X))#butlast (map snd (half_segments_of_aform X))" by (rule nth_equalityI) (auto simp: nth_Cons hd_conv_nth nth_butlast subsequent_half_segments split: nat.split) lemma le_less_Suc_eq: "x - Suc 0 \ (i::nat) \ i < x \ x - Suc 0 = i" by simp lemma map_snd_half_segments_aux_eq: "half_segments_of_aform X \ [] \ map snd (half_segments_of_aform X) = tl (map fst (half_segments_of_aform X)) @ [snd (last (half_segments_of_aform X))]" by (rule nth_equalityI) (auto simp: nth_Cons hd_conv_nth nth_append nth_tl subsequent_half_segments not_less last_conv_nth algebra_simps dest!: le_less_Suc_eq split: nat.split) lemma ccw'_sortedP_snd_half_segments_of_aform: "ccw'.sortedP (lowest_vertex (fst X, nlex_pdevs (snd X))) (map snd (half_segments_of_aform X))" by (auto simp: half_segments_of_aform_def Let_def intro!: ccw'.sortedP.Cons ccw'_sortedP_polychain_of_snd ccw'_sortedP_scaled_inl) lemma lex_half_segments_lowest_vertex: assumes "(c, d) \ set (half_segments_of_aform X)" shows "lex d (lowest_vertex (fst X, nlex_pdevs (snd X)))" unfolding half_segments_of_aform_def Let_def by (rule lex_polychain_of_center[OF assms[unfolded half_segments_of_aform_def Let_def], unfolded snd_conv]) (auto simp: list_all_iff lex_def dest!: in_set_inl_lex) lemma lex_half_segments_lowest_vertex': assumes "d \ set (map snd (half_segments_of_aform X))" shows "lex d (lowest_vertex (fst X, nlex_pdevs (snd X)))" using assms by (auto intro: lex_half_segments_lowest_vertex) lemma lex_half_segments_last: assumes "(c, d) \ set (half_segments_of_aform X)" shows "lex (snd (last (half_segments_of_aform X))) d" using assms unfolding half_segments_of_aform_def Let_def by (rule lex_polychain_of_last) (auto simp: list_all_iff lex_def dest!: in_set_inl_lex) lemma lex_half_segments_last': assumes "d \ set (map snd (half_segments_of_aform X))" shows "lex (snd (last (half_segments_of_aform X))) d" using assms by (auto intro: lex_half_segments_last) lemma ccw'_half_segments_lowest_last: assumes set_butlast: "(c, d) \ set (butlast (half_segments_of_aform X))" assumes ne: "inl (snd X) \ []" shows "ccw' (lowest_vertex (fst X, nlex_pdevs (snd X))) d (snd (last (half_segments_of_aform X)))" using set_butlast unfolding half_segments_of_aform_def Let_def by (rule ccw'_polychain_of_sorted_center_last) (auto simp: ne ccw'_sortedP_scaled_inl) lemma distinct_fst_half_segments: "distinct (map fst (half_segments_of_aform X))" by (auto simp: half_segments_of_aform_def list_all_iff lex_scale1_zero simp del: scaleR_Pair intro!: distinct_fst_polychain_of dest: in_set_inl_nonzero in_set_inl_lex) lemma distinct_snd_half_segments: "distinct (map snd (half_segments_of_aform X))" by (auto simp: half_segments_of_aform_def list_all_iff lex_scale1_zero simp del: scaleR_Pair intro!: distinct_snd_polychain_of dest: in_set_inl_nonzero in_set_inl_lex) subsection \Mirror\ definition "mirror_point x y = 2 *\<^sub>R x - y" lemma ccw'_mirror_point3[simp]: "ccw' (mirror_point x y) (mirror_point x z) (mirror_point x w) \ ccw' y z w " by (auto simp: mirror_point_def det3_def' ccw'_def algebra_simps) lemma mirror_point_self_inverse[simp]: fixes x::"'a::real_vector" shows "mirror_point p (mirror_point p x) = x" by (auto simp: mirror_point_def scaleR_2) lemma mirror_half_segments_of_aform: assumes "e \ UNIV \ {-1 <..< 1}" assumes "length (half_segments_of_aform X) \ 1" shows "list_all (\seg. ccw' (fst seg) (snd seg) (aform_val e X)) (map (pairself (mirror_point (fst X))) (half_segments_of_aform X))" unfolding list_all_length proof safe let ?m = "map (pairself (mirror_point (fst X))) (half_segments_of_aform X)" fix n assume "n < length ?m" hence "ccw' (fst (half_segments_of_aform X ! n)) (snd (half_segments_of_aform X ! n)) (aform_val (- e) X)" using assms by (auto dest!: nth_mem intro!: half_segments_of_aform_strict) also have "aform_val (-e) X = 2 *\<^sub>R fst X - aform_val e X" by (auto simp: aform_val_def pdevs_val_sum algebra_simps scaleR_2 sum_negf) finally have le: "ccw' (fst (half_segments_of_aform X ! n)) (snd (half_segments_of_aform X ! n)) (2 *\<^sub>R fst X - aform_val e X)" . have eq: "(map (pairself (mirror_point (fst X))) (half_segments_of_aform X) ! n) = (2 *\<^sub>R fst X - fst ((half_segments_of_aform X) ! n), 2 *\<^sub>R fst X - snd ((half_segments_of_aform X) ! n))" using \n < length ?m\ by (cases "half_segments_of_aform X ! n") (auto simp add: mirror_point_def) show "ccw' (fst (?m ! n)) (snd (?m ! n)) (aform_val e X)" using le unfolding eq by (auto simp: algebra_simps ccw'_def det3_def') qed lemma last_half_segments: assumes "half_segments_of_aform X \ []" shows "snd (last (half_segments_of_aform X)) = mirror_point (fst X) (lowest_vertex (fst X, nlex_pdevs (snd X)))" using assms by (auto simp add: half_segments_of_aform_def Let_def lowest_vertex_def mirror_point_def scaleR_2 scaleR_sum_list[symmetric] last_polychain_of sum_list_distinct_selsort inl_def sum_list_independent_pdevs sum_list_list_of_pdevs) lemma convex_polychain_map_mirror: assumes "convex_polychain hs" shows "convex_polychain (map (pairself (mirror_point x)) hs)" proof (rule convex_polychainI) qed (insert assms, auto simp: convex_polychain_def polychain_map_pairself pairself_apply mirror_point_def det3_def' algebra_simps) lemma ccw'_mirror_point0: "ccw' (mirror_point x y) z w \ ccw' y (mirror_point x z) (mirror_point x w)" by (metis ccw'_mirror_point3 mirror_point_self_inverse) lemma ccw'_sortedP_mirror: "ccw'.sortedP x0 (map (mirror_point p0) xs) \ ccw'.sortedP (mirror_point p0 x0) xs" by (induct xs) (simp_all add: linorder_list0.sortedP.Nil linorder_list0.sortedP_Cons_iff ccw'_mirror_point0) lemma ccw'_sortedP_mirror2: "ccw'.sortedP (mirror_point p0 x0) (map (mirror_point p0) xs) \ ccw'.sortedP x0 xs" using ccw'_sortedP_mirror[of "mirror_point p0 x0" p0 xs] by simp lemma map_mirror_o_snd_polychain_of_eq: "map (mirror_point x0 \ snd) (polychain_of y xs) = map snd (polychain_of (mirror_point x0 y) (map uminus xs))" by (induct xs arbitrary: x0 y) (auto simp: mirror_point_def o_def algebra_simps) lemma lowest_vertex_eq_mirror_last: "half_segments_of_aform X \ [] \ (lowest_vertex (fst X, nlex_pdevs (snd X))) = mirror_point (fst X) (snd (last (half_segments_of_aform X)))" using last_half_segments[of X] by simp lemma snd_last: "xs \ [] \ snd (last xs) = last (map snd xs)" by (induct xs) auto lemma mirror_point_snd_last_eq_lowest: "half_segments_of_aform X \ [] \ mirror_point (fst X) (last (map snd (half_segments_of_aform X))) = lowest_vertex (fst X, nlex_pdevs (snd X))" by (simp add: lowest_vertex_eq_mirror_last snd_last) lemma lex_mirror_point: "lex (mirror_point x0 a) (mirror_point x0 b) \ lex b a" by (auto simp: mirror_point_def lex_def) lemma ccw'_mirror_point: "ccw' (mirror_point x0 a) (mirror_point x0 b) (mirror_point x0 c) \ ccw' a b c" by (auto simp: mirror_point_def) lemma inj_mirror_point: "inj (mirror_point (x::'a::real_vector))" by (auto simp: mirror_point_def inj_on_def algebra_simps) lemma distinct_map_mirror_point_eq: "distinct (map (mirror_point (x::'a::real_vector)) xs) = distinct xs" by (auto simp: distinct_map intro!: subset_inj_on[OF inj_mirror_point]) lemma eq_self_mirror_iff: fixes x::"'a::real_vector" shows "x = mirror_point y x \ x = y" by (auto simp: mirror_point_def algebra_simps scaleR_2[symmetric]) subsection \Full Segments\ definition segments_of_aform::"point aform \ (point * point) list" where "segments_of_aform X = (let hs = half_segments_of_aform X in hs @ map (pairself (mirror_point (fst X))) hs)" lemma segments_of_aform_strict: assumes "e \ UNIV \ {-1 <..< 1}" assumes "length (half_segments_of_aform X) \ 1" shows "list_all (\seg. ccw' (fst seg) (snd seg) (aform_val e X)) (segments_of_aform X)" using assms by (auto simp: segments_of_aform_def Let_def mirror_half_segments_of_aform half_segments_of_aform_strict_all) lemma mirror_point_aform_val[simp]: "mirror_point (fst X) (aform_val e X) = aform_val (- e) X" by (auto simp: mirror_point_def aform_val_def pdevs_val_sum algebra_simps scaleR_2 sum_negf) lemma in_set_segments_of_aform_aform_valE: assumes "(x2, y2) \ set (segments_of_aform X)" obtains e where "y2 = aform_val e X" "e \ UNIV \ {-1 .. 1}" using assms proof (auto simp: segments_of_aform_def Let_def) assume "(x2, y2) \ set (half_segments_of_aform X)" from in_set_half_segments_of_aform_aform_valE[OF this] obtain e where "y2 = aform_val e X" "e \ UNIV \ {- 1..1}" by auto thus ?thesis .. next fix a b aa ba assume "((a, b), aa, ba) \ set (half_segments_of_aform X)" from in_set_half_segments_of_aform_aform_valE[OF this] obtain e where e: "(aa, ba) = aform_val e X" "e \ UNIV \ {- 1..1}" by auto assume "y2 = mirror_point (fst X) (aa, ba)" hence "y2 = aform_val (-e) X" "(-e) \ UNIV \ {-1 .. 1}" using e by auto thus ?thesis .. qed lemma last_half_segments_eq_mirror_hd: assumes "half_segments_of_aform X \ []" shows "snd (last (half_segments_of_aform X)) = mirror_point (fst X) (fst (hd (half_segments_of_aform X)))" by (simp add: last_half_segments assms fst_hd_half_segments_of_aform) lemma polychain_segments_of_aform: "polychain (segments_of_aform X)" by (auto simp: segments_of_aform_def Let_def polychain_half_segments_of_aform polychain_map_pairself last_half_segments_eq_mirror_hd hd_map pairself_apply intro!: polychain_appendI) lemma segments_of_aform_closed: assumes "segments_of_aform X \ []" shows "fst (hd (segments_of_aform X)) = snd (last (segments_of_aform X))" using assms by (auto simp: segments_of_aform_def Let_def fst_hd_half_segments_of_aform last_map pairself_apply last_half_segments mirror_point_def) lemma convex_polychain_segments_of_aform: assumes "1 < length (half_segments_of_aform X)" shows "convex_polychain (segments_of_aform X)" unfolding segments_of_aform_def Let_def using ccw_hd_last_half_segments_dirvec[OF assms] by (intro convex_polychain_appendI) (auto simp: convex_polychain_half_segments_of_aform convex_polychain_map_mirror dirvec_minus hd_map pairself_apply last_half_segments mirror_point_def fst_hd_half_segments_of_aform det3_def' algebra_simps ccw'_def intro!: polychain_appendI polychain_half_segments_of_aform polychain_map_pairself) lemma convex_polygon_segments_of_aform: assumes "1 < length (half_segments_of_aform X)" shows "convex_polygon (segments_of_aform X)" proof - from assms have hne: "half_segments_of_aform X \ []" by auto from convex_polychain_segments_of_aform[OF assms] have "convex_polychain (segments_of_aform X)" . thus ?thesis by (auto simp: convex_polygon_def segments_of_aform_closed) qed lemma previous_segments_of_aformE: assumes "(y, z) \ set (segments_of_aform X)" obtains x where "(x, y) \ set (segments_of_aform X)" proof - from assms have ne[simp]: "segments_of_aform X \ []" by auto from assms obtain i where i: "i set (segments_of_aform X)" by (metis fst_conv hd_conv_nth i(2) last_in_set ne snd_conv surj_pair) thus ?thesis .. next case (Suc j) have "(fst (segments_of_aform X ! j), snd (segments_of_aform X ! j)) \ set (segments_of_aform X)" using Suc i(1) by auto also from i have "y = fst (segments_of_aform X ! i)" by auto hence "snd (segments_of_aform X ! j) = y" using polychain_segments_of_aform[of X] i(1) Suc by (auto simp: polychain_def Suc) finally have "(fst (segments_of_aform X ! j), y) \ set (segments_of_aform X)" . thus ?thesis .. qed qed lemma fst_compose_pairself: "fst o pairself f = f o fst" and snd_compose_pairself: "snd o pairself f = f o snd" by (auto simp: pairself_apply) lemma in_set_butlastI: "xs \ [] \ x \ set xs \ x \ last xs \ x \ set (butlast xs)" by (induct xs) (auto split: if_splits) lemma distinct_in_set_butlastD: "x \ set (butlast xs) \ distinct xs \ x \ last xs" by (induct xs) auto lemma distinct_in_set_tlD: "x \ set (tl xs) \ distinct xs \ x \ hd xs" by (induct xs) auto lemma ccw'_sortedP_snd_segments_of_aform: assumes "length (inl (snd X)) > 1" shows "ccw'.sortedP (lowest_vertex (fst X, nlex_pdevs (snd X))) (butlast (map snd (segments_of_aform X)))" proof cases assume "[] = half_segments_of_aform X" from this show ?thesis by (simp add: segments_of_aform_def Let_def ccw'.sortedP.Nil) next assume H: "[] \ half_segments_of_aform X" let ?m = "mirror_point (fst X)" have ne: "inl (snd X) \ []" using assms by auto have "ccw'.sortedP (lowest_vertex (fst X, nlex_pdevs (snd X))) (map snd (half_segments_of_aform X) @ butlast (map (?m \ snd) (half_segments_of_aform X)))" proof (rule ccw'.sortedP_appendI) show "ccw'.sortedP (lowest_vertex (fst X, nlex_pdevs (snd X))) (map snd (half_segments_of_aform X))" by (rule ccw'_sortedP_snd_half_segments_of_aform) have "butlast (map (?m \ snd) (half_segments_of_aform X)) = butlast (map (?m \ snd) (polychain_of (lowest_vertex (fst X, nlex_pdevs (snd X))) (map ((*\<^sub>R) 2) (ccw.selsort 0 (inl (snd X))))))" by (simp add: half_segments_of_aform_def) also have "\ = map snd (butlast (polychain_of (?m (lowest_vertex (fst X, nlex_pdevs (snd X)))) (map uminus (map ((*\<^sub>R) 2) (ccw.selsort 0 (inl (snd X)))))))" (is "_ = map snd (butlast (polychain_of ?x ?xs))") by (simp add: map_mirror_o_snd_polychain_of_eq map_butlast) also { have "ccw'.sortedP 0 ?xs" by (intro ccw'_sortedP_uminus ccw'_sortedP_scaled_inl) moreover have "ccw'.sortedP ?x (map snd (polychain_of ?x ?xs))" unfolding ccw'_sortedP_mirror[symmetric] map_map map_mirror_o_snd_polychain_of_eq by (auto simp add: o_def intro!: ccw'_sortedP_polychain_of_snd ccw'_sortedP_scaled_inl) ultimately have "ccw'.sortedP (snd (last (polychain_of ?x ?xs))) (map snd (butlast (polychain_of ?x ?xs)))" by (rule ccw'_sortedP_convex_rotate_aux) } also have "(snd (last (polychain_of ?x ?xs))) = ?m (last (map snd (half_segments_of_aform X)))" by (simp add: half_segments_of_aform_def ne map_mirror_o_snd_polychain_of_eq last_map[symmetric, where f="?m"] last_map[symmetric, where f="snd"]) also have "\ = lowest_vertex (fst X, nlex_pdevs (snd X))" using ne H by (auto simp: lowest_vertex_eq_mirror_last snd_last) finally show "ccw'.sortedP (lowest_vertex (fst X, nlex_pdevs (snd X))) (butlast (map (?m \ snd) (half_segments_of_aform X)))" . next fix x y assume seg: "x \ set (map snd (half_segments_of_aform X))" and mseg: "y \ set (butlast (map (?m \ snd) (half_segments_of_aform X)))" from seg assms have neq_Nil: "inl (snd X) \ []" "half_segments_of_aform X \ []" by auto from seg obtain a where a: "(a, x) \ set (half_segments_of_aform X)" by auto from mseg obtain b where mirror_y: "(b, ?m y) \ set (butlast (half_segments_of_aform X))" by (auto simp: map_butlast[symmetric]) let ?l = "lowest_vertex (fst X, nlex_pdevs (snd X))" let ?ml = "?m ?l" have mirror_eq_last: "?ml = snd (last (half_segments_of_aform X))" using seg H by (intro last_half_segments[symmetric]) simp define r where "r = ?l + (0, abs (snd x - snd ?l) + abs (snd y - snd ?l) + abs (snd ?ml - snd ?l) + 1)" have d1: "x \ r" "y \ r" "?l \ r" "?ml \ r" by (auto simp: r_def plus_prod_def prod_eq_iff) have "distinct (map (?m \ snd) (half_segments_of_aform X))" unfolding map_comp_map[symmetric] unfolding o_def distinct_map_mirror_point_eq by (rule distinct_snd_half_segments) from distinct_in_set_butlastD[OF \y \ _\ this] have "?l \ y" by (simp add: neq_Nil lowest_vertex_eq_mirror_last last_map) moreover have "?l \ ?ml" using neq_Nil by (auto simp add: eq_self_mirror_iff lowest_vertex_eq_center_iff inl_def) ultimately have d2: "?l \ y" "?l \ ?ml" by auto have *: "ccw' ?l (?m y) ?ml" by (metis mirror_eq_last ccw'_half_segments_lowest_last mirror_y neq_Nil(1)) have "ccw' ?ml y ?l" by (rule ccw'_mirror_point[of "fst X"]) (simp add: *) hence lmy: "ccw' ?l ?ml y" by (simp add: ccw'_def det3_def' algebra_simps) let ?ccw = "ccw' (lowest_vertex (fst X, nlex_pdevs (snd X))) x y" { assume "x \ ?ml" hence x_butlast: "(a, x) \ set (butlast (half_segments_of_aform X))" unfolding mirror_eq_last using a by (auto intro!: in_set_butlastI simp: prod_eq_iff) have "ccw' ?l x ?ml" by (metis mirror_eq_last ccw'_half_segments_lowest_last x_butlast neq_Nil(1)) } note lxml = this { assume "x = ?ml" hence ?ccw by (simp add: lmy) } moreover { assume "x \ ?ml" "y = ?ml" hence ?ccw by (simp add: lxml) } moreover { assume d3: "x \ ?ml" "y \ ?ml" from \x \ set _\ have "x \ set (map snd (half_segments_of_aform X))" by force hence "x \ set (tl (map fst (half_segments_of_aform X)))" using d3 unfolding map_snd_half_segments_aux_eq[OF neq_Nil(2)] by (auto simp: mirror_eq_last) from distinct_in_set_tlD[OF this distinct_fst_half_segments] have "?l \ x" by (simp add: fst_hd_half_segments_of_aform neq_Nil hd_map) from lxml[OF \x \ ?ml\] \ccw' ?l ?ml y\ have d4: "x \ y" by (rule neq_left_right_of lxml) have "distinct5 x ?ml y r ?l" using d1 d2 \?l \ x\ d3 d4 by simp_all moreover note _ moreover have "lex x ?l" by (rule lex_half_segments_lowest_vertex) fact hence "ccw ?l r x" unfolding r_def by (rule lex_ccw_left) simp moreover have "lex ?ml ?l" using last_in_set[OF H[symmetric]] by (auto simp: mirror_eq_last intro: lex_half_segments_lowest_vertex') hence "ccw ?l r ?ml" unfolding r_def by (rule lex_ccw_left) simp moreover have "lex (?m (lowest_vertex (fst X, nlex_pdevs (snd X)))) (?m y)" using mirror_y by (force dest!: in_set_butlastD intro: lex_half_segments_last' simp: mirror_eq_last ) hence "lex y ?l" by (rule lex_mirror_point) hence "ccw ?l r y" unfolding r_def by (rule lex_ccw_left) simp moreover from \x \ ?ml\ have "ccw ?l x ?ml" by (simp add: ccw_def lxml) moreover from lmy have "ccw ?l ?ml y" by (simp add: ccw_def) ultimately have "ccw ?l x y" by (rule ccw.transitive[where S=UNIV]) simp moreover { have "x \ ?l" using \?l \ x\ by simp moreover have "lex (?m y) (?m ?ml)" using mirror_y by (force intro: lex_half_segments_lowest_vertex in_set_butlastD) hence "lex ?ml y" by (rule lex_mirror_point) moreover from a have "lex ?ml x" unfolding mirror_eq_last by (rule lex_half_segments_last) moreover note \lex y ?l\ \lex x ?l\ \ccw' ?l x ?ml\ \ccw' ?l ?ml y\ ultimately have ncoll: "\ coll ?l x y" by (rule not_coll_ordered_lexI) } ultimately have ?ccw by (simp add: ccw_def) } ultimately show ?ccw by blast qed thus ?thesis using H by (simp add: segments_of_aform_def Let_def butlast_append snd_compose_pairself) qed lemma polychain_of_segments_of_aform1: assumes "length (segments_of_aform X) = 1" shows "False" using assms by (auto simp: segments_of_aform_def Let_def half_segments_of_aform_def add_is_1 split: if_split_asm) lemma polychain_of_segments_of_aform2: assumes "segments_of_aform X = [x, y]" assumes "between (fst x, snd x) p" shows "p \ convex hull set (map fst (segments_of_aform X))" proof - from polychain_segments_of_aform[of X] segments_of_aform_closed[of X] assms have "fst y = snd x" "snd y = fst x" by (simp_all add: polychain_def) thus ?thesis using assms by (cases x) (auto simp: between_mem_convex_hull) qed lemma append_eq_2: assumes "length xs = length ys" shows "xs @ ys = [x, y] \ xs = [x] \ ys = [y]" using assms proof (cases xs) case (Cons z zs) thus ?thesis using assms by (cases zs) auto qed simp lemma segments_of_aform_line_segment: assumes "segments_of_aform X = [x, y]" assumes "e \ UNIV \ {-1 .. 1}" shows "aform_val e X \ closed_segment (fst x) (snd x)" proof - from pdevs_val_pdevs_of_list_inl2E[OF \e \ _\, of "snd X"] obtain e' where e': "pdevs_val e (snd X) = pdevs_val e' (pdevs_of_list (inl (snd X)))" "e' \ UNIV \ {- 1..1}" . from e' have "0 \ 1 + e' 0" by (auto simp: Pi_iff dest!: spec[where x=0]) with assms e' show ?thesis by (auto simp: segments_of_aform_def Let_def append_eq_2 half_segments_of_aform_def polychain_of_singleton_iff mirror_point_def ccw.selsort_singleton_iff lowest_vertex_def aform_val_def sum_list_nlex_eq_sum_list_inl closed_segment_def Pi_iff intro!: exI[where x="(1 + e' 0) / 2"]) (auto simp: algebra_simps) qed subsection \Continuous Generalization\ lemma LIMSEQ_minus_fract_mult: "(\n. r * (1 - 1 / real (Suc (Suc n)))) \ r" by (rule tendsto_eq_rhs[OF tendsto_mult[where a=r and b = 1]]) (auto simp: inverse_eq_divide[symmetric] simp del: of_nat_Suc intro: filterlim_compose[OF LIMSEQ_inverse_real_of_nat filterlim_Suc] tendsto_eq_intros) lemma det3_nonneg_segments_of_aform: assumes "e \ UNIV \ {-1 .. 1}" assumes "length (half_segments_of_aform X) \ 1" shows "list_all (\seg. det3 (fst seg) (snd seg) (aform_val e X) \ 0) (segments_of_aform X)" unfolding list_all_iff proof safe fix a b c d assume seg: "((a, b), c, d) \ set (segments_of_aform X)" (is "?seg \ _") define normal_of_segment where "normal_of_segment = (\((a0, a1), b0, b1). (b1 - a1, a0 - b0)::real*real)" define support_of_segment where "support_of_segment = (\(a, b). normal_of_segment (a, b) \ a)" have "closed ((\x. x \ normal_of_segment ?seg) -` {..support_of_segment ?seg})" (is "closed ?cl") by (auto intro!: continuous_intros closed_vimage) moreover define f where "f n i = e i * ( 1 - 1 / (Suc (Suc n)))" for n i have "\n. aform_val (f n) X \ ?cl" proof fix n have "f n \ UNIV \ {-1 <..< 1}" using assms by (auto simp: f_def Pi_iff intro!: less_one_multI minus_one_less_multI) from list_allD[OF segments_of_aform_strict[OF this assms(2)] seg] show "aform_val (f n) X \ (\x. x \ normal_of_segment ((a, b), c, d)) -` {..support_of_segment ((a, b), c, d)}" by (auto simp: list_all_iff normal_of_segment_def support_of_segment_def det3_def' field_simps inner_prod_def ccw'_def) qed moreover have "\i. (\n. f n i) \ e i" unfolding f_def by (rule LIMSEQ_minus_fract_mult) hence "(\n. aform_val (f n) X) \ aform_val e X" by (auto simp: aform_val_def pdevs_val_sum intro!: tendsto_intros) ultimately have "aform_val e X \ ?cl" by (rule closed_sequentially) thus "det3 (fst ?seg) (snd ?seg) (aform_val e X) \ 0" by (auto simp: list_all_iff normal_of_segment_def support_of_segment_def det3_def' field_simps inner_prod_def) qed lemma det3_nonneg_segments_of_aformI: assumes "e \ UNIV \ {-1 .. 1}" assumes "length (half_segments_of_aform X) \ 1" assumes "seg \ set (segments_of_aform X)" shows "det3 (fst seg) (snd seg) (aform_val e X) \ 0" using assms det3_nonneg_segments_of_aform by (auto simp: list_all_iff) subsection \Intersection of Vertical Line with Segment\ fun intersect_segment_xline'::"nat \ point * point \ real \ point option" where "intersect_segment_xline' p ((x0, y0), (x1, y1)) xl = (if x0 \ xl \ xl \ x1 then if x0 = x1 then Some ((min y0 y1), (max y0 y1)) else let yl = truncate_down p (truncate_down p (real_divl p (y1 - y0) (x1 - x0) * (xl - x0)) + y0); yr = truncate_up p (truncate_up p (real_divr p (y1 - y0) (x1 - x0) * (xl - x0)) + y0) in Some (yl, yr) else None)" lemma norm_pair_fst0[simp]: "norm (0, x) = norm x" by (auto simp: norm_prod_def) lemmas add_right_mono_le = order_trans[OF add_right_mono] lemmas mult_right_mono_le = order_trans[OF mult_right_mono] lemmas add_right_mono_ge = order_trans[OF _ add_right_mono] lemmas mult_right_mono_ge = order_trans[OF _ mult_right_mono] lemma dest_segment: fixes x b::real assumes "(x, b) \ closed_segment (x0, y0) (x1, y1)" assumes "x0 \ x1" shows "b = (y1 - y0) * (x - x0) / (x1 - x0) + y0" proof - from assms obtain u where u: "x = x0 *\<^sub>R (1 - u) + u * x1" "b = y0 *\<^sub>R (1 - u) + u * y1" "0 \ u" "u \ 1" by (auto simp: closed_segment_def algebra_simps) show "b = (y1 - y0) * (x - x0) / (x1 - x0) + y0 " using assms by (auto simp: closed_segment_def field_simps u) qed lemma intersect_segment_xline': assumes "intersect_segment_xline' prec (p0, p1) x = Some (m, M)" shows "closed_segment p0 p1 \ {p. fst p = x} \ {(x, m) .. (x, M)}" using assms proof (cases p0) case (Pair x0 y0) note p0 = this show ?thesis proof (cases p1) case (Pair x1 y1) note p1 = this { assume "x0 = x1" "x = x1" "m = min y0 y1" "M = max y0 y1" hence ?thesis by (force simp: abs_le_iff p0 p1 min_def max_def algebra_simps dest: segment_bound split: if_split_asm) } thus ?thesis using assms by (auto simp: abs_le_iff p0 p1 split: if_split_asm intro!: truncate_up_le truncate_down_le add_right_mono_le[OF truncate_down] add_right_mono_le[OF real_divl] add_right_mono_le[OF mult_right_mono_le[OF real_divl]] add_right_mono_ge[OF _ truncate_up] add_right_mono_ge[OF _ mult_right_mono_ge[OF _ real_divr]] dest!: dest_segment) qed qed lemma in_segment_fst_le: fixes x0 x1 b::real assumes "x0 \ x1" "(x, b) \ closed_segment (x0, y0) (x1, y1)" shows "x \ x1" using assms using mult_left_mono[OF \x0 \ x1\, where c="1 - u" for u] by (force simp add: min_def max_def split: if_split_asm simp add: algebra_simps not_le closed_segment_def) lemma in_segment_fst_ge: fixes x0 x1 b::real assumes "x0 \ x1" "(x, b) \ closed_segment (x0, y0) (x1, y1)" shows "x0 \ x" using assms using mult_left_mono[OF \x0 \ x1\] by (force simp add: min_def max_def split: if_split_asm simp add: algebra_simps not_le closed_segment_def) lemma intersect_segment_xline'_eq_None: assumes "intersect_segment_xline' prec (p0, p1) x = None" assumes "fst p0 \ fst p1" shows "closed_segment p0 p1 \ {p. fst p = x} = {}" using assms by (cases p0, cases p1) (auto simp: abs_le_iff split: if_split_asm dest: in_segment_fst_le in_segment_fst_ge) fun intersect_segment_xline where "intersect_segment_xline prec ((a, b), (c, d)) x = (if a \ c then intersect_segment_xline' prec ((a, b), (c, d)) x else intersect_segment_xline' prec ((c, d), (a, b)) x)" lemma closed_segment_commute: "closed_segment a b = closed_segment b a" by (meson convex_contains_segment convex_closed_segment dual_order.antisym ends_in_segment) lemma intersect_segment_xline: assumes "intersect_segment_xline prec (p0, p1) x = Some (m, M)" shows "closed_segment p0 p1 \ {p. fst p = x} \ {(x, m) .. (x, M)}" using assms by (cases p0, cases p1) (auto simp: closed_segment_commute split: if_split_asm simp del: intersect_segment_xline'.simps dest!: intersect_segment_xline') lemma intersect_segment_xline_fst_snd: assumes "intersect_segment_xline prec seg x = Some (m, M)" shows "closed_segment (fst seg) (snd seg) \ {p. fst p = x} \ {(x, m) .. (x, M)}" using intersect_segment_xline[of prec "fst seg" "snd seg" x m M] assms by simp lemma intersect_segment_xline_eq_None: assumes "intersect_segment_xline prec (p0, p1) x = None" shows "closed_segment p0 p1 \ {p. fst p = x} = {}" using assms by (cases p0, cases p1) (auto simp: closed_segment_commute split: if_split_asm simp del: intersect_segment_xline'.simps dest!: intersect_segment_xline'_eq_None) lemma inter_image_empty_iff: "(X \ {p. f p = x} = {}) \ (x \ f ` X)" by auto lemma fst_closed_segment[simp]: "fst ` closed_segment a b = closed_segment (fst a) (fst b)" by (force simp: closed_segment_def) lemma intersect_segment_xline_eq_empty: fixes p0 p1::"real * real" assumes "closed_segment p0 p1 \ {p. fst p = x} = {}" shows "intersect_segment_xline prec (p0, p1) x = None" using assms by (cases p0, cases p1) (auto simp: inter_image_empty_iff closed_segment_eq_real_ivl split: if_split_asm) lemma intersect_segment_xline_le: assumes "intersect_segment_xline prec y xl = Some (m0, M0)" shows "m0 \ M0" using assms by (cases y) (auto simp: min_def split: if_split_asm intro!: truncate_up_le truncate_down_le order_trans[OF real_divl] order_trans[OF _ real_divr] mult_right_mono) lemma intersect_segment_xline_None_iff: fixes p0 p1::"real * real" shows "intersect_segment_xline prec (p0, p1) x = None \ closed_segment p0 p1 \ {p. fst p = x} = {}" by (auto intro!: intersect_segment_xline_eq_empty dest!: intersect_segment_xline_eq_None) subsection \Bounds on Vertical Intersection with Oriented List of Segments\ primrec bound_intersect_2d where "bound_intersect_2d prec [] x = None" | "bound_intersect_2d prec (X # Xs) xl = (case bound_intersect_2d prec Xs xl of None \ intersect_segment_xline prec X xl | Some (m, M) \ (case intersect_segment_xline prec X xl of None \ Some (m, M) | Some (m', M') \ Some (min m' m, max M' M)))" lemma bound_intersect_2d_eq_None: assumes "bound_intersect_2d prec Xs x = None" assumes "X \ set Xs" shows "intersect_segment_xline prec X x = None" using assms by (induct Xs) (auto split: option.split_asm) lemma bound_intersect_2d_upper: assumes "bound_intersect_2d prec Xs x = Some (m, M)" obtains X m' where "X \ set Xs" "intersect_segment_xline prec X x = Some (m', M)" "\X m' M' . X \ set Xs \ intersect_segment_xline prec X x = Some (m', M') \ M' \ M" proof atomize_elim show "\X m'. X \ set Xs \ intersect_segment_xline prec X x = Some (m', M) \ (\X m' M'. X \ set Xs \ intersect_segment_xline prec X x = Some (m', M') \ M' \ M)" using assms proof (induct Xs arbitrary: m M) case Nil thus ?case by (simp add: bound_intersect_2d_def) next case (Cons X Xs) show ?case proof (cases "bound_intersect_2d prec Xs x") case None thus ?thesis using Cons by (intro exI[where x=X] exI[where x=m]) (simp del: intersect_segment_xline.simps add: bound_intersect_2d_eq_None) next case (Some mM) note Some1=this then obtain m' M' where mM: "mM = (m', M')" by (cases mM) from Cons(1)[OF Some[unfolded mM]] obtain X' m'' where X': "X' \ set Xs" and m'': "intersect_segment_xline prec X' x = Some (m'', M')" and max: "\X m' M'a. X \ set Xs \ intersect_segment_xline prec X x = Some (m', M'a) \ M'a \ M'" by auto show ?thesis proof (cases "intersect_segment_xline prec X x") case None thus ?thesis using Some1 mM Cons(2) X' m'' max by (intro exI[where x= X'] exI[where x= m'']) (auto simp del: intersect_segment_xline.simps split: option.split_asm) next case (Some mM''') thus ?thesis using Some1 mM Cons(2) X' m'' by (cases mM''') (force simp: max_def min_def simp del: intersect_segment_xline.simps split: option.split_asm if_split_asm dest!: max intro!: exI[where x= "if M' \ snd mM''' then X' else X"] exI[where x= "if M' \ snd mM''' then m'' else fst mM'''"]) qed qed qed qed lemma bound_intersect_2d_lower: assumes "bound_intersect_2d prec Xs x = Some (m, M)" obtains X M' where "X \ set Xs" "intersect_segment_xline prec X x = Some (m, M')" "\X m' M' . X \ set Xs \ intersect_segment_xline prec X x = Some (m', M') \ m \ m'" proof atomize_elim show "\X M'. X \ set Xs \ intersect_segment_xline prec X x = Some (m, M') \ (\X m' M'. X \ set Xs \ intersect_segment_xline prec X x = Some (m', M') \ m \ m')" using assms proof (induct Xs arbitrary: m M) case Nil thus ?case by (simp add: bound_intersect_2d_def) next case (Cons X Xs) show ?case proof (cases "bound_intersect_2d prec Xs x") case None thus ?thesis using Cons by (intro exI[where x= X]) (simp del: intersect_segment_xline.simps add: bound_intersect_2d_eq_None) next case (Some mM) note Some1=this then obtain m' M' where mM: "mM = (m', M')" by (cases mM) from Cons(1)[OF Some[unfolded mM]] obtain X' M'' where X': "X' \ set Xs" and M'': "intersect_segment_xline prec X' x = Some (m', M'')" and min: "\X m'a M'. X \ set Xs \ intersect_segment_xline prec X x = Some (m'a, M') \ m' \ m'a" by auto show ?thesis proof (cases "intersect_segment_xline prec X x") case None thus ?thesis using Some1 mM Cons(2) X' M'' min by (intro exI[where x= X'] exI[where x= M'']) (auto simp del: intersect_segment_xline.simps split: option.split_asm) next case (Some mM''') thus ?thesis using Some1 mM Cons(2) X' M'' min by (cases mM''') (force simp: max_def min_def simp del: intersect_segment_xline.simps split: option.split_asm if_split_asm dest!: min intro!: exI[where x= "if m' \ fst mM''' then X' else X"] exI[where x= "if m' \ fst mM''' then M'' else snd mM'''"]) qed qed qed qed lemma bound_intersect_2d: assumes "bound_intersect_2d prec Xs x = Some (m, M)" shows "(\(p1, p2) \ set Xs. closed_segment p1 p2) \ {p. fst p = x} \ {(x, m) .. (x, M)}" proof (clarsimp, safe) fix b x0 y0 x1 y1 assume H: "((x0, y0), x1, y1) \ set Xs" "(x, b) \ closed_segment (x0, y0) (x1, y1)" hence "intersect_segment_xline prec ((x0, y0), x1, y1) x \ None" by (intro notI) (auto dest!: intersect_segment_xline_eq_None simp del: intersect_segment_xline.simps) then obtain e f where ef: "intersect_segment_xline prec ((x0, y0), x1, y1) x = Some (e, f)" by auto with H have "m \ e" by (auto intro: bound_intersect_2d_lower[OF assms]) also have "\ \ b" using intersect_segment_xline[OF ef] H by force finally show "m \ b" . have "b \ f" using intersect_segment_xline[OF ef] H by force also have "\ \ M" using H ef by (auto intro: bound_intersect_2d_upper[OF assms]) finally show "b \ M" . qed lemma bound_intersect_2d_eq_None_iff: "bound_intersect_2d prec Xs x = None \ (\X\set Xs. intersect_segment_xline prec X x = None)" by (induct Xs) (auto simp: split: option.split_asm) lemma bound_intersect_2d_nonempty: assumes "bound_intersect_2d prec Xs x = Some (m, M)" shows "(\(p1, p2) \ set Xs. closed_segment p1 p2) \ {p. fst p = x} \ {}" proof - from assms have "bound_intersect_2d prec Xs x \ None" by simp then obtain p1 p2 where "(p1, p2) \ set Xs" "intersect_segment_xline prec (p1, p2) x \ None" unfolding bound_intersect_2d_eq_None_iff by auto hence "closed_segment p1 p2 \ {p. fst p = x} \ {}" by (simp add: intersect_segment_xline_None_iff) thus ?thesis using \(p1, p2) \ set Xs\ by auto qed lemma bound_intersect_2d_le: assumes "bound_intersect_2d prec Xs x = Some (m, M)" shows "m \ M" proof - from bound_intersect_2d_nonempty[OF assms] bound_intersect_2d[OF assms] show "m \ M" by auto qed subsection \Bounds on Vertical Intersection with General List of Segments\ definition "bound_intersect_2d_ud prec X xl = (case segments_of_aform X of [] \ if fst (fst X) = xl then let m = snd (fst X) in Some (m, m) else None | [x, y] \ intersect_segment_xline prec x xl | xs \ (case bound_intersect_2d prec (filter (\((x1, y1), x2, y2). x1 < x2) xs) xl of Some (m, M') \ (case bound_intersect_2d prec (filter (\((x1, y1), x2, y2). x1 > x2) xs) xl of Some (m', M) \ Some (min m m', max M M') | None \ None) | None \ None))" lemma list_cases4: "\x P. (x = [] \ P) \ (\y. x = [y] \ P) \ (\y z. x = [y, z] \ P) \ (\w y z zs. x = w # y # z # zs \ P) \ P" by (metis list.exhaust) lemma bound_intersect_2d_ud_segments_of_aform_le: "bound_intersect_2d_ud prec X xl = Some (m0, M0) \ m0 \ M0" by (cases "segments_of_aform X" rule: list_cases4) (auto simp: Let_def bound_intersect_2d_ud_def min_def max_def intersect_segment_xline_le if_split_eq1 split: option.split_asm prod.split_asm list.split_asm dest!: bound_intersect_2d_le) lemma pdevs_domain_eq_empty_iff[simp]: "pdevs_domain (snd X) = {} \ snd X = zero_pdevs" by (auto simp: intro!: pdevs_eqI) lemma ccw_contr_on_line_left: assumes "det3 (a, b) (x, c) (x, d) \ 0" "a > x" shows "d \ c" proof - from assms have "d * (a - x) \ c * (a - x)" by (auto simp: det3_def' algebra_simps) with assms show "c \ d" by auto qed lemma ccw_contr_on_line_right: assumes "det3 (a, b) (x, c) (x, d) \ 0" "a < x" shows "d \ c" proof - from assms have "c * (x - a) \ d * (x - a)" by (auto simp: det3_def' algebra_simps) with assms show "d \ c" by auto qed lemma singleton_contrE: assumes "\x y. x \ y \ x \ X \ y \ X \ False" assumes "X \ {}" obtains x where "X = {x}" using assms by blast lemma segment_intersection_singleton: fixes xl and a b::"real * real" defines "i \ closed_segment a b \ {p. fst p = xl}" assumes ne1: "fst a \ fst b" assumes upper: "i \ {}" obtains p1 where "i = {p1}" proof (rule singleton_contrE[OF _ upper]) fix x y assume H: "x \ y" "x \ i" "y \ i" then obtain u v where uv: "x = u *\<^sub>R b + (1 - u) *\<^sub>R a" "y = v *\<^sub>R b + (1 - v) *\<^sub>R a" "0 \ u" "u \ 1" "0 \ v" "v \ 1" - by (auto simp: closed_segment_def i_def field_simps) - hence "u = v * (fst a - fst b) / (fst a - fst b)" using ne1 H(2,3) - by (auto simp: closed_segment_def field_simps i_def) - hence "u = v" by (simp add: ne1) - thus False using H uv by simp + by (auto simp add: closed_segment_def i_def field_simps) + then have "x + u *\<^sub>R a = a + u *\<^sub>R b" "y + v *\<^sub>R a = a + v *\<^sub>R b" + by simp_all + then have "fst (x + u *\<^sub>R a) = fst (a + u *\<^sub>R b)" "fst (y + v *\<^sub>R a) = fst (a + v *\<^sub>R b)" + by simp_all + then have "u = v * (fst a - fst b) / (fst a - fst b)" + using ne1 H(2,3) \0 \ u\ \u \ 1\ \0 \ v\ \v \ 1\ + by (simp add: closed_segment_def i_def field_simps) + then have "u = v" + by (simp add: ne1) + then show False using H uv + by simp qed lemma bound_intersect_2d_ud_segments_of_aform: assumes "bound_intersect_2d_ud prec X xl = Some (m0, M0)" assumes "e \ UNIV \ {-1 .. 1}" shows "{aform_val e X} \ {x. fst x = xl} \ {(xl, m0) .. (xl, M0)}" proof safe fix a b assume safeassms: "(a, b) = aform_val e X" "xl = fst (a, b)" hence fst_aform_val: "fst (aform_val e X) = xl" by simp { assume len: "length (segments_of_aform X) > 2" with assms obtain m M m' M' where lb: "bound_intersect_2d prec [((x1, y1), x2, y2)\segments_of_aform X . x1 < x2] xl = Some (m, M')" and ub: "bound_intersect_2d prec [((x1, y1), x2, y2)\segments_of_aform X . x2 < x1] xl = Some (m', M)" and minmax: "m0 = min m m'" "M0 = max M M'" by (auto simp: bound_intersect_2d_ud_def split: option.split_asm list.split_asm ) from bound_intersect_2d_upper[OF ub] obtain X1 m1 where upper: "X1 \ set [((x1, y1), x2, y2)\segments_of_aform X . x2 < x1]" "intersect_segment_xline prec X1 xl = Some (m1, M)" by metis from bound_intersect_2d_lower[OF lb] obtain X2 M2 where lower: "X2 \ set [((x1, y1), x2, y2)\segments_of_aform X . x1 < x2]" "intersect_segment_xline prec X2 xl = Some (m, M2)" by metis from upper(1) lower(1) have X1: "X1 \ set (segments_of_aform X)" "fst (fst X1) > fst (snd X1)" and X2: "X2 \ set (segments_of_aform X)" "fst (fst X2) < fst (snd X2)" by auto note upper_seg = intersect_segment_xline_fst_snd[OF upper(2)] note lower_seg = intersect_segment_xline_fst_snd[OF lower(2)] from len have lh: "length (half_segments_of_aform X) \ 1" by (auto simp: segments_of_aform_def Let_def) from X1 have ne1: "fst (fst X1) \ fst (snd X1)" by simp moreover have "closed_segment (fst X1) (snd X1) \ {p. fst p = xl} \ {}" using upper(2) by (simp add: intersect_segment_xline_None_iff[of prec, symmetric]) ultimately obtain p1 where p1: "closed_segment (fst X1) (snd X1) \ {p. fst p = xl} = {p1}" by (rule segment_intersection_singleton) then obtain u where u: "p1 = (1 - u) *\<^sub>R fst X1 + u *\<^sub>R (snd X1)" "0 \ u" "u \ 1" by (auto simp: closed_segment_def algebra_simps) have coll1: "det3 (fst X1) p1 (aform_val e X) \ 0" and coll1': "det3 p1 (snd X1) (aform_val e X) \ 0" unfolding atomize_conj using u by (cases "u = 0 \ u = 1") (auto simp: u(1) intro: det3_nonneg_scaleR_segment1 det3_nonneg_scaleR_segment2 det3_nonneg_segments_of_aformI[OF \e \ _\ lh X1(1)]) from X2 have ne2: "fst (fst X2) \ fst (snd X2)" by simp moreover have "closed_segment (fst X2) (snd X2) \ {p. fst p = xl} \ {}" using lower(2) by (simp add: intersect_segment_xline_None_iff[of prec, symmetric]) ultimately obtain p2 where p2: "closed_segment (fst X2) (snd X2) \ {p. fst p = xl} = {p2}" by (rule segment_intersection_singleton) then obtain v where v: "p2 = (1 - v) *\<^sub>R fst X2 + v *\<^sub>R (snd X2)" "0 \ v" "v \ 1" by (auto simp: closed_segment_def algebra_simps) have coll2: "det3 (fst X2) p2 (aform_val e X) \ 0" and coll2': "det3 p2 (snd X2) (aform_val e X) \ 0" unfolding atomize_conj using v by (cases "v = 0 \ v = 1") (auto simp: v(1) intro: det3_nonneg_scaleR_segment1 det3_nonneg_scaleR_segment2 det3_nonneg_segments_of_aformI[OF \e \ _\ lh X2(1)]) from in_set_segments_of_aform_aform_valE [of "fst X1" "snd X1" X, unfolded prod.collapse, OF X1(1)] obtain e1s where e1s: "snd X1 = aform_val e1s X" "e1s \ UNIV \ {- 1..1}" . from previous_segments_of_aformE [of "fst X1" "snd X1" X, unfolded prod.collapse, OF X1(1)] obtain fX0 where "(fX0, fst X1) \ set (segments_of_aform X)" . from in_set_segments_of_aform_aform_valE[OF this] obtain e1f where e1f: "fst X1 = aform_val e1f X" "e1f \ UNIV \ {- 1..1}" . have "p1 \ closed_segment (aform_val e1f X) (aform_val e1s X)" using p1 by (auto simp: e1s e1f) with segment_in_aform_val[OF e1s(2) e1f(2), of X] obtain ep1 where ep1: "ep1 \ UNIV \ {-1 .. 1}" "p1 = aform_val ep1 X" by (auto simp: Affine_def valuate_def closed_segment_commute) from in_set_segments_of_aform_aform_valE [of "fst X2" "snd X2" X, unfolded prod.collapse, OF X2(1)] obtain e2s where e2s: "snd X2 = aform_val e2s X" "e2s \ UNIV \ {- 1..1}" . from previous_segments_of_aformE [of "fst X2" "snd X2" X, unfolded prod.collapse, OF X2(1)] obtain fX02 where "(fX02, fst X2) \ set (segments_of_aform X)" . from in_set_segments_of_aform_aform_valE[OF this] obtain e2f where e2f: "fst X2 = aform_val e2f X" "e2f \ UNIV \ {- 1..1}" . have "p2 \ closed_segment (aform_val e2f X) (aform_val e2s X)" using p2 by (auto simp: e2s e2f) with segment_in_aform_val[OF e2f(2) e2s(2), of X] obtain ep2 where ep2: "ep2 \ UNIV \ {-1 .. 1}" "p2 = aform_val ep2 X" by (auto simp: Affine_def valuate_def) from det3_nonneg_segments_of_aformI[OF ep2(1), of X "(fst X1, snd X1)", unfolded prod.collapse, OF lh X1(1), unfolded ep2(2)[symmetric]] have c2: "det3 (fst X1) (snd X1) p2 \ 0" . hence c12: "det3 (fst X1) p1 p2 \ 0" using u by (cases "u = 0") (auto simp: u(1) det3_nonneg_scaleR_segment2) from det3_nonneg_segments_of_aformI[OF ep1(1), of X "(fst X2, snd X2)", unfolded prod.collapse, OF lh X2(1), unfolded ep1(2)[symmetric]] have c1: "det3 (fst X2) (snd X2) p1 \ 0" . hence c21: "det3 (fst X2) p2 p1 \ 0" using v by (cases "v = 0") (auto simp: v(1) det3_nonneg_scaleR_segment2) from p1 p2 have p1p2xl: "fst p1 = xl" "fst p2 = xl" by (auto simp: det3_def') from upper_seg p1 have "snd p1 \ M" by (auto simp: less_eq_prod_def) from lower_seg p2 have "m \ snd p2" by (auto simp: less_eq_prod_def) { have *: "(fst p1, snd (aform_val e X)) = aform_val e X" by (simp add: prod_eq_iff p1p2xl fst_aform_val) hence coll: "det3 (fst (fst X1), snd (fst X1)) (fst p1, snd p1) (fst p1, snd (aform_val e X)) \ 0" and coll': "det3 (fst (snd X1), snd (snd X1)) (fst p1, snd (aform_val e X)) (fst p1, snd p1) \ 0" using coll1 coll1' by (auto simp: det3_rotate) have "snd (aform_val e X) \ M" proof (cases "fst (fst X1) = xl") case False have "fst (fst X1) \ fst p1" unfolding u using X1 by (auto simp: algebra_simps intro!: mult_left_mono u) moreover have "fst (fst X1) \ fst p1" using False by (simp add: p1p2xl) ultimately have "fst (fst X1) > fst p1" by simp from ccw_contr_on_line_left[OF coll this] show ?thesis using \snd p1 \ M\ by simp next case True have "fst (snd X1) * (1 - u) \ fst (fst X1) * (1 - u)" using X1 u by (auto simp: intro!: mult_right_mono) hence "fst (snd X1) \ fst p1" unfolding u by (auto simp: algebra_simps) moreover have "fst (snd X1) \ fst p1" using True ne1 by (simp add: p1p2xl) ultimately have "fst (snd X1) < fst p1" by simp from ccw_contr_on_line_right[OF coll' this] show ?thesis using \snd p1 \ M\ by simp qed } moreover { have "(fst p2, snd (aform_val e X)) = aform_val e X" by (simp add: prod_eq_iff p1p2xl fst_aform_val) hence coll: "det3 (fst (fst X2), snd (fst X2)) (fst p2, snd p2) (fst p2, snd (aform_val e X)) \ 0" and coll': "det3 (fst (snd X2), snd (snd X2)) (fst p2, snd (aform_val e X)) (fst p2, snd p2) \ 0" using coll2 coll2' by (auto simp: det3_rotate) have "m \ snd (aform_val e X)" proof (cases "fst (fst X2) = xl") case False have "fst (fst X2) \ fst p2" unfolding v using X2 by (auto simp: algebra_simps intro!: mult_left_mono v) moreover have "fst (fst X2) \ fst p2" using False by (simp add: p1p2xl) ultimately have "fst (fst X2) < fst p2" by simp from ccw_contr_on_line_right[OF coll this] show ?thesis using \m \ snd p2\ by simp next case True have "(1 - v) * fst (snd X2) \ (1 - v) * fst (fst X2)" using X2 v by (auto simp: intro!: mult_left_mono) hence "fst (snd X2) \ fst p2" unfolding v by (auto simp: algebra_simps) moreover have "fst (snd X2) \ fst p2" using True ne2 by (simp add: p1p2xl) ultimately have "fst (snd X2) > fst p2" by simp from ccw_contr_on_line_left[OF coll' this] show ?thesis using \m \ snd p2\ by simp qed } ultimately have "aform_val e X \ {(xl, m) .. (xl, M)}" by (auto simp: less_eq_prod_def fst_aform_val) hence "aform_val e X \ {(xl, m0) .. (xl, M0)}" by (auto simp: minmax less_eq_prod_def) } moreover { assume "length (segments_of_aform X) = 2" then obtain a b where s: "segments_of_aform X = [a, b]" by (auto simp: numeral_2_eq_2 length_Suc_conv) from segments_of_aform_line_segment[OF this assms(2)] have "aform_val e X \ closed_segment (fst a) (snd a)" . moreover from assms have "intersect_segment_xline prec a xl = Some (m0, M0)" by (auto simp: bound_intersect_2d_ud_def s) note intersect_segment_xline_fst_snd[OF this] ultimately have "aform_val e X \ {(xl, m0) .. (xl, M0)}" by (auto simp: less_eq_prod_def fst_aform_val) } moreover { assume "length (segments_of_aform X) = 1" from polychain_of_segments_of_aform1[OF this] have "aform_val e X \ {(xl, m0) .. (xl, M0)}" by auto } moreover { assume len: "length (segments_of_aform X) = 0" hence "independent_pdevs (map snd (list_of_pdevs (nlex_pdevs (snd X)))) = []" by (simp add: segments_of_aform_def Let_def half_segments_of_aform_def inl_def) hence "snd X = zero_pdevs" by (subst (asm) independent_pdevs_eq_Nil_iff) (auto simp: list_all_iff list_of_pdevs_def) hence "aform_val e X = fst X" by (simp add: aform_val_def) with len assms have "aform_val e X \ {(xl, m0) .. (xl, M0)}" by (auto simp: bound_intersect_2d_ud_def Let_def split: if_split_asm) } ultimately have "aform_val e X \ {(xl, m0)..(xl, M0)}" by arith thus "(a, b) \ {(fst (a, b), m0)..(fst (a, b), M0)}" using safeassms by simp qed subsection \Approximation from Orthogonal Directions\ definition inter_aform_plane_ortho:: "nat \ 'a::executable_euclidean_space aform \ 'a \ real \ 'a aform option" where "inter_aform_plane_ortho p Z n g = do { mMs \ those (map (\b. bound_intersect_2d_ud p (inner2_aform Z n b) g) Basis_list); let l = (\(b, m)\zip Basis_list (map fst mMs). m *\<^sub>R b); let u = (\(b, M)\zip Basis_list (map snd mMs). M *\<^sub>R b); Some (aform_of_ivl l u) }" lemma those_eq_SomeD: assumes "those (map f xs) = Some ys" shows "ys = map (the o f) xs \ (\i.\y. i < length xs \ f (xs ! i) = Some y)" using assms by (induct xs arbitrary: ys) (auto split: option.split_asm simp: o_def nth_Cons split: nat.split) lemma sum_list_zip_map: assumes "distinct xs" shows "(\(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 inter_aform_plane_ortho_overappr: assumes "inter_aform_plane_ortho p Z n g = Some X" shows "{x. \i \ Basis. x \ i \ {y. (g, y) \ (\x. (x \ n, x \ i)) ` Affine Z}} \ Affine X" proof - let ?inter = "(\b. bound_intersect_2d_ud p (inner2_aform Z n b) g)" obtain xs where xs: "those (map ?inter Basis_list) = Some xs" using assms by (cases "those (map ?inter Basis_list)") (auto simp: inter_aform_plane_ortho_def) from those_eq_SomeD[OF this] obtain y' where xs_eq: "xs = map (the \ ?inter) Basis_list" and y': "\i. i < length (Basis_list::'a list) \ ?inter (Basis_list ! i) = Some (y' i)" by metis have "\(i::'a) \ Basis. \ji::'a. i\Basis \ j i < length (Basis_list::'a list)" "\i::'a. i\Basis \ i = Basis_list ! j i" by metis define y where "y = y' o j" with y' j have y: "\i. i \ Basis \ ?inter i = Some (y i)" by (metis comp_def) hence y_le: "\i. i \ Basis \ fst (y i) \ snd (y i)" by (auto intro!: bound_intersect_2d_ud_segments_of_aform_le) hence "(\b\Basis. fst (y b) *\<^sub>R b) \ (\b\Basis. snd (y b) *\<^sub>R b)" by (auto simp: eucl_le[where 'a='a]) with assms have X: "Affine X = {\b\Basis. fst (y b) *\<^sub>R b..\b\Basis. snd (y b) *\<^sub>R b}" by (auto simp: inter_aform_plane_ortho_def sum_list_zip_map xs xs_eq y Affine_aform_of_ivl) show ?thesis proof safe fix x assume x: "\i\Basis. x \ i \ {y. (g, y) \ (\x. (x \ n, x \ i)) ` Affine Z}" { fix i::'a assume i: "i \ Basis" from x i have x_in2: "(g, x \ i) \ (\x. (x \ n, x \ i)) ` Affine Z" by auto from x_in2 obtain e where e: "e \ UNIV \ {- 1..1}" and g: "g = aform_val e Z \ n" and x: "x \ i = aform_val e Z \ i" by (auto simp: Affine_def valuate_def) have "{aform_val e (inner2_aform Z n i)} = {aform_val e (inner2_aform Z n i)} \ {x. fst x = g}" by (auto simp: g inner2_def) also from y[OF \i \ Basis\] have "?inter i = Some (fst (y i), snd (y i))" by simp note bound_intersect_2d_ud_segments_of_aform[OF this e] finally have "x \ i \ {fst (y i) .. snd (y i)}" by (auto simp: x inner2_def) } thus "x \ Affine X" unfolding X by (auto simp: eucl_le[where 'a='a]) qed qed lemma inter_proj_eq: fixes n g l defines "G \ {x. x \ n = g}" shows "(\x. x \ l) ` (Z \ G) = {y. (g, y) \ (\x. (x \ n, x \ l)) ` Z}" by (auto simp: G_def) lemma inter_overappr: fixes n \ l shows "Z \ {x. x \ n = g} \ {x. \i \ Basis. x \ i \ {y. (g, y) \ (\x. (x \ n, x \ i)) ` Z}}" by auto lemma inter_inter_aform_plane_ortho: assumes "inter_aform_plane_ortho p Z n g = Some X" shows "Affine Z \ {x. x \ n = g} \ Affine X" proof - note inter_overappr[of "Affine Z" n g] also note inter_aform_plane_ortho_overappr[OF assms] finally show ?thesis . qed subsection \``Completeness'' of Intersection\ abbreviation "encompasses x seg \ det3 (fst seg) (snd seg) x \ 0" lemma encompasses_cases: "encompasses x seg \ encompasses x (snd seg, fst seg)" by (auto simp: det3_def' algebra_simps) lemma list_all_encompasses_cases: assumes "list_all (encompasses p) (x # y # zs)" obtains "list_all (encompasses p) [x, y, (snd y, fst x)]" | "list_all (encompasses p) ((fst x, snd y)#zs)" using encompasses_cases proof assume "encompasses p (snd y, fst x)" hence "list_all (encompasses p) [x, y, (snd y, fst x)]" using assms by (auto simp: list_all_iff) thus ?thesis .. next assume "encompasses p (snd (snd y, fst x), fst (snd y, fst x))" hence "list_all (encompasses p) ((fst x, snd y)#zs)" using assms by (auto simp: list_all_iff) thus ?thesis .. qed lemma triangle_encompassing_polychain_of: assumes "det3 p a b \ 0" "det3 p b c \ 0" "det3 p c a \ 0" assumes "ccw' a b c" shows "p \ convex hull {a, b, c}" proof - from assms have nn: "det3 b c p \ 0" "det3 c a p \ 0" "det3 a b p \ 0" "det3 a b c \ 0" by (auto simp: det3_def' algebra_simps) have "det3 a b c *\<^sub>R p = det3 b c p *\<^sub>R a + det3 c a p *\<^sub>R b + det3 a b p *\<^sub>R c" by (auto simp: det3_def' algebra_simps prod_eq_iff) hence "inverse (det3 a b c) *\<^sub>R (det3 a b c *\<^sub>R p) = inverse (det3 a b c) *\<^sub>R (det3 b c p *\<^sub>R a + det3 c a p *\<^sub>R b + det3 a b p *\<^sub>R c)" by simp with assms have p_eq: "p = (det3 b c p / det3 a b c) *\<^sub>R a + (det3 c a p / det3 a b c) *\<^sub>R b + (det3 a b p / det3 a b c) *\<^sub>R c" (is "_ = scaleR ?u _ + scaleR ?v _ + scaleR ?w _") by (simp add: inverse_eq_divide algebra_simps ccw'_def) have det_eq: "det3 b c p / det3 a b c + det3 c a p / det3 a b c + det3 a b p / det3 a b c = 1" using assms(4) by (simp add: add_divide_distrib[symmetric] det3_def' algebra_simps ccw'_def) show ?thesis unfolding convex_hull_3 using assms(4) by (blast intro: exI[where x="?u"] exI[where x="?v"] exI[where x="?w"] intro!: p_eq divide_nonneg_nonneg nn det_eq) qed lemma encompasses_convex_polygon3: assumes "list_all (encompasses p) (x#y#z#zs)" assumes "convex_polygon (x#y#z#zs)" assumes "ccw'.sortedP (fst x) (map snd (butlast (x#y#z#zs)))" shows "p \ convex hull (set (map fst (x#y#z#zs)))" using assms proof (induct zs arbitrary: x y z p) case Nil thus ?case by (auto simp: det3_def' algebra_simps elim!: ccw'.sortedP_Cons ccw'.sortedP_Nil intro!: triangle_encompassing_polychain_of) next case (Cons w ws) from Cons.prems(2) have "snd y = fst z" by auto from Cons.prems(1) show ?case proof (rule list_all_encompasses_cases) assume "list_all (encompasses p) [x, y, (snd y, fst x)]" hence "p \ convex hull {fst x, fst y, snd y}" using Cons.prems by (auto simp: det3_def' algebra_simps elim!: ccw'.sortedP_Cons ccw'.sortedP_Nil intro!: triangle_encompassing_polychain_of) thus ?case by (rule rev_subsetD[OF _ hull_mono]) (auto simp: \snd y = fst z\) next assume *: "list_all (encompasses p) ((fst x, snd y) # z # w # ws)" from Cons.prems have enc: "ws \ [] \ encompasses p (last ws)" by (auto simp: list_all_iff) have "set (map fst ((fst x, snd y)#z#w#ws)) \ set (map fst (x # y # z # w # ws))" by auto moreover { note * moreover have "convex_polygon ((fst x, snd y) # z # w # ws)" by (metis convex_polygon_skip Cons.prems(2,3)) moreover have "ccw'.sortedP (fst (fst x, snd y)) (map snd (butlast ((fst x, snd y) # z # w # ws)))" using Cons.prems(3) by (auto elim!: ccw'.sortedP_Cons intro!: ccw'.sortedP.Cons ccw'.sortedP.Nil split: if_split_asm) ultimately have "p \ convex hull set (map fst ((fst x, snd y)#z#w#ws))" by (rule Cons.hyps) } ultimately show "p \ convex hull set (map fst (x # y # z # w # ws))" by (rule subsetD[OF hull_mono]) qed qed lemma segments_of_aform_empty_Affine_eq: assumes "segments_of_aform X = []" shows "Affine X = {fst X}" proof - have "independent_pdevs (map snd (list_of_pdevs (nlex_pdevs (snd X)))) = [] \ (list_of_pdevs (nlex_pdevs (snd X))) = []" by (subst independent_pdevs_eq_Nil_iff) (auto simp: list_all_iff list_of_pdevs_def ) with assms show ?thesis by (force simp: aform_val_def list_of_pdevs_def Affine_def valuate_def segments_of_aform_def Let_def half_segments_of_aform_def inl_def) qed lemma not_segments_of_aform_singleton: "segments_of_aform X \ [x]" by (auto simp: segments_of_aform_def Let_def add_is_1 dest!: arg_cong[where f=length]) lemma encompasses_segments_of_aform_in_AffineI: assumes "length (segments_of_aform X) > 2" assumes "list_all (encompasses p) (segments_of_aform X)" shows "p \ Affine X" proof - from assms(1) obtain x y z zs where eq: "segments_of_aform X = x#y#z#zs" by (cases "segments_of_aform X" rule: list_cases4) auto hence "fst x = fst (hd (half_segments_of_aform X))" by (metis segments_of_aform_def hd_append list.map_disc_iff list.sel(1)) also have "\ = lowest_vertex (fst X, nlex_pdevs (snd X))" using assms by (intro fst_hd_half_segments_of_aform) (auto simp: segments_of_aform_def) finally have fstx: "fst x = lowest_vertex (fst X, nlex_pdevs (snd X))" . have "p \ convex hull (set (map fst (segments_of_aform X)))" using assms(2) unfolding eq proof (rule encompasses_convex_polygon3) show "convex_polygon (x # y # z # zs)" using assms(1) unfolding eq[symmetric] by (intro convex_polygon_segments_of_aform) (simp add: segments_of_aform_def Let_def) show "ccw'.sortedP (fst x) (map snd (butlast (x # y # z # zs)))" using assms(1) unfolding fstx map_butlast eq[symmetric] by (intro ccw'_sortedP_snd_segments_of_aform) (simp add: segments_of_aform_def Let_def half_segments_of_aform_def) qed also have "\ \ convex hull (Affine X)" proof (rule hull_mono, safe) fix a b assume "(a, b) \ set (map fst (segments_of_aform X))" then obtain c d where "((a, b), c, d) \ set (segments_of_aform X)" by auto from previous_segments_of_aformE[OF this] obtain x where "(x, a, b) \ set (segments_of_aform X)" by auto from in_set_segments_of_aform_aform_valE[OF this] obtain e where "(a, b) = aform_val e X" "e \ UNIV \ {- 1..1}" by auto thus "(a, b) \ Affine X" by (auto simp: Affine_def valuate_def image_iff) qed also have "\ = Affine X" by (simp add: convex_Affine convex_hull_eq) finally show ?thesis . qed end diff --git a/thys/Prime_Number_Theorem/Prime_Number_Theorem_Library.thy b/thys/Prime_Number_Theorem/Prime_Number_Theorem_Library.thy --- a/thys/Prime_Number_Theorem/Prime_Number_Theorem_Library.thy +++ b/thys/Prime_Number_Theorem/Prime_Number_Theorem_Library.thy @@ -1,2041 +1,2041 @@ section \Auxiliary material\ theory Prime_Number_Theorem_Library imports Zeta_Function.Zeta_Function "HOL-Real_Asymp.Real_Asymp" begin (* TODO: Move *) lemma asymp_equivD_strong: assumes "f \[F] g" "eventually (\x. f x \ 0 \ g x \ 0) F" shows "((\x. f x / g x) \ 1) F" proof - from assms(1) have "((\x. if f x = 0 \ g x = 0 then 1 else f x / g x) \ 1) F" by (rule asymp_equivD) also have "?this \ ?thesis" by (intro filterlim_cong eventually_mono[OF assms(2)]) auto finally show ?thesis . qed lemma frontier_real_Ici [simp]: fixes a :: real shows "frontier {a..} = {a}" unfolding frontier_def by (auto simp: interior_real_atLeast) lemma sum_upto_ln_conv_sum_upto_mangoldt: "sum_upto (\n. ln (real n)) x = sum_upto (\n. mangoldt n * nat \x / real n\) x" proof - have "sum_upto (\n. ln (real n)) x = sum_upto (\n. \d | d dvd n. mangoldt d) x" by (intro sum_upto_cong) (simp_all add: mangoldt_sum) also have "\ = sum_upto (\k. sum_upto (\d. mangoldt k) (x / real k)) x" by (rule sum_upto_sum_divisors) also have "\ = sum_upto (\n. mangoldt n * nat \x / real n\) x" unfolding sum_upto_altdef by (simp add: mult_ac) finally show ?thesis . qed lemma ln_fact_conv_sum_upto_mangoldt: "ln (fact n) = sum_upto (\k. mangoldt k * (n div k)) n" proof - have [simp]: "{0<..Suc n} = insert (Suc n) {0<..n}" for n by auto have "ln (fact n) = sum_upto (\n. ln (real n)) n" by (induction n) (auto simp: sum_upto_altdef nat_add_distrib ln_mult) also have "\ = sum_upto (\k. mangoldt k * (n div k)) n" unfolding sum_upto_ln_conv_sum_upto_mangoldt by (intro sum_upto_cong) (auto simp: floor_divide_of_nat_eq) finally show ?thesis . qed lemma powr_sum: "x \ 0 \ finite A \ x powr sum f A = (\y\A. x powr f y)" by (simp add: powr_def exp_sum sum_distrib_right) lemma fds_abs_converges_comparison_test: fixes s :: "'a :: dirichlet_series" assumes "eventually (\n. norm (fds_nth f n) \ fds_nth g n) at_top" and "fds_converges g (s \ 1)" shows "fds_abs_converges f s" unfolding fds_abs_converges_def proof (rule summable_comparison_test_ev) from assms(2) show "summable (\n. fds_nth g n / n powr (s \ 1))" by (auto simp: fds_converges_def) from assms(1) eventually_gt_at_top[of 0] show "eventually (\n. norm (norm (fds_nth f n / nat_power n s)) \ fds_nth g n / real n powr (s \ 1)) at_top" by eventually_elim (auto simp: norm_divide norm_nat_power intro!: divide_right_mono) qed lemma fds_converges_scaleR [intro]: assumes "fds_converges f s" shows "fds_converges (c *\<^sub>R f) s" proof - from assms have "summable (\n. c *\<^sub>R (fds_nth f n / nat_power n s))" by (intro summable_scaleR_right) (auto simp: fds_converges_def) also have "(\n. c *\<^sub>R (fds_nth f n / nat_power n s)) = (\n. (c *\<^sub>R fds_nth f n / nat_power n s))" by (simp add: scaleR_conv_of_real) finally show ?thesis by (simp add: fds_converges_def) qed lemma fds_abs_converges_scaleR [intro]: assumes "fds_abs_converges f s" shows "fds_abs_converges (c *\<^sub>R f) s" proof - from assms have "summable (\n. abs c * norm (fds_nth f n / nat_power n s))" by (intro summable_mult) (auto simp: fds_abs_converges_def) also have "(\n. abs c * norm (fds_nth f n / nat_power n s)) = (\n. norm ((c *\<^sub>R fds_nth f n) / nat_power n s))" by (simp add: norm_divide) finally show ?thesis by (simp add: fds_abs_converges_def) qed lemma conv_abscissa_scaleR: "conv_abscissa (scaleR c f) \ conv_abscissa f" by (rule conv_abscissa_mono) auto lemma abs_conv_abscissa_scaleR: "abs_conv_abscissa (scaleR c f) \ abs_conv_abscissa f" by (rule abs_conv_abscissa_mono) auto lemma fds_converges_mult_const_left [intro]: "fds_converges f s \ fds_converges (fds_const c * f) s" by (auto simp: fds_converges_def dest: summable_mult[of _ c]) lemma fds_abs_converges_mult_const_left [intro]: "fds_abs_converges f s \ fds_abs_converges (fds_const c * f) s" by (auto simp: fds_abs_converges_def norm_mult norm_divide dest: summable_mult[of _ "norm c"]) lemma conv_abscissa_mult_const_left: "conv_abscissa (fds_const c * f) \ conv_abscissa f" by (intro conv_abscissa_mono) auto lemma abs_conv_abscissa_mult_const_left: "abs_conv_abscissa (fds_const c * f) \ abs_conv_abscissa f" by (intro abs_conv_abscissa_mono) auto lemma fds_converges_mult_const_right [intro]: "fds_converges f s \ fds_converges (f * fds_const c) s" by (auto simp: fds_converges_def dest: summable_mult2[of _ c]) lemma fds_abs_converges_mult_const_right [intro]: "fds_abs_converges f s \ fds_abs_converges (f * fds_const c) s" by (auto simp: fds_abs_converges_def norm_mult norm_divide dest: summable_mult2[of _ "norm c"]) lemma conv_abscissa_mult_const_right: "conv_abscissa (f * fds_const c) \ conv_abscissa f" by (intro conv_abscissa_mono) auto lemma abs_conv_abscissa_mult_const_right: "abs_conv_abscissa (f * fds_const c) \ abs_conv_abscissa f" by (intro abs_conv_abscissa_mono) auto lemma bounded_coeffs_imp_fds_abs_converges: fixes s :: "'a :: dirichlet_series" and f :: "'a fds" assumes "Bseq (fds_nth f)" "s \ 1 > 1" shows "fds_abs_converges f s" proof - from assms obtain C where C: "\n. norm (fds_nth f n) \ C" by (auto simp: Bseq_def) show ?thesis proof (rule fds_abs_converges_comparison_test) from \s \ 1 > 1\ show "fds_converges (C *\<^sub>R fds_zeta) (s \ 1)" by (intro fds_abs_converges_imp_converges) auto from C show "eventually (\n. norm (fds_nth f n) \ fds_nth (C *\<^sub>R fds_zeta) n) at_top" by (intro always_eventually) (auto simp: fds_nth_zeta) qed qed lemma bounded_coeffs_imp_fds_abs_converges': fixes s :: "'a :: dirichlet_series" and f :: "'a fds" assumes "Bseq (\n. fds_nth f n * nat_power n s0)" "s \ 1 > 1 - s0 \ 1" shows "fds_abs_converges f s" proof - have "fds_nth (fds_shift s0 f) = (\n. fds_nth f n * nat_power n s0)" by (auto simp: fun_eq_iff) with assms have "Bseq (fds_nth (fds_shift s0 f))" by simp with assms(2) have "fds_abs_converges (fds_shift s0 f) (s + s0)" by (intro bounded_coeffs_imp_fds_abs_converges) (auto simp: algebra_simps) thus ?thesis by simp qed lemma bounded_coeffs_imp_abs_conv_abscissa_le: fixes s :: "'a :: dirichlet_series" and f :: "'a fds" and c :: ereal assumes "Bseq (\n. fds_nth f n * nat_power n s)" "1 - s \ 1 \ c" shows "abs_conv_abscissa f \ c" proof (rule abs_conv_abscissa_leI_weak) fix x assume "c < ereal x" have "ereal (1 - s \ 1) \ c" by fact also have "\ < ereal x" by fact finally have "1 - s \ 1 < ereal x" by simp thus "fds_abs_converges f (of_real x)" by (intro bounded_coeffs_imp_fds_abs_converges'[OF assms(1)]) auto qed lemma bounded_coeffs_imp_abs_conv_abscissa_le_1: fixes s :: "'a :: dirichlet_series" and f :: "'a fds" assumes "Bseq (\n. fds_nth f n)" shows "abs_conv_abscissa f \ 1" proof - have [simp]: "fds_nth f n * nat_power n 0 = fds_nth f n" for n by (cases "n = 0") auto show ?thesis by (rule bounded_coeffs_imp_abs_conv_abscissa_le[where s = 0]) (insert assms, auto simp:) qed (* TODO: replace library version *) (* EXAMPLE: This might make a good example to illustrate real_asymp *) lemma fixes a b c :: real assumes ab: "a + b > 0" and c: "c < -1" shows set_integrable_powr_at_top: "(\x. (b + x) powr c) absolutely_integrable_on {a<..}" and set_lebesgue_integral_powr_at_top: "(\x\{a<..}. ((b + x) powr c) \lborel) = -((b + a) powr (c + 1) / (c + 1))" and powr_has_integral_at_top: "((\x. (b + x) powr c) has_integral -((b + a) powr (c + 1) / (c + 1))) {a<..}" proof - let ?f = "\x. (b + x) powr c" and ?F = "\x. (b + x) powr (c + 1) / (c + 1)" have limits: "((?F \ real_of_ereal) \ ?F a) (at_right (ereal a))" "((?F \ real_of_ereal) \ 0) (at_left \)" using c ab unfolding ereal_tendsto_simps1 by (real_asymp simp: field_simps)+ have 1: "set_integrable lborel (einterval a \) ?f" using ab c limits by (intro interval_integral_FTC_nonneg) (auto intro!: derivative_eq_intros) thus 2: "?f absolutely_integrable_on {a<..}" by (auto simp: set_integrable_def integrable_completion) have "LBINT x=ereal a..\. (b + x) powr c = 0 - ?F a" using ab c limits by (intro interval_integral_FTC_nonneg) (auto intro!: derivative_eq_intros) thus 3: "(\x\{a<..}. ((b + x) powr c) \lborel) = -((b + a) powr (c + 1) / (c + 1))" by (simp add: interval_integral_to_infinity_eq) show "(?f has_integral -((b + a) powr (c + 1) / (c + 1))) {a<..}" using set_borel_integral_eq_integral[OF 1] 3 by (simp add: has_integral_iff) qed lemma fds_converges_altdef2: "fds_converges f s \ convergent (\N. eval_fds (fds_truncate N f) s)" unfolding fds_converges_def summable_iff_convergent' eval_fds_truncate by (auto simp: not_le intro!: convergent_cong always_eventually sum.mono_neutral_right) lemma tendsto_eval_fds_truncate: assumes "fds_converges f s" shows "(\N. eval_fds (fds_truncate N f) s) \ eval_fds f s" proof - have "(\N. eval_fds (fds_truncate N f) s) \ eval_fds f s \ (\N. \i\N. fds_nth f i / nat_power i s) \ eval_fds f s" unfolding eval_fds_truncate by (intro filterlim_cong always_eventually allI sum.mono_neutral_left) (auto simp: not_le) also have \ using assms by (simp add: fds_converges_iff sums_def' atLeast0AtMost) finally show ?thesis . qed lemma linepath_translate_left: "linepath (c + a) (c + a) = (\x. c + a) \ linepath a b" by (auto simp: fun_eq_iff linepath_def algebra_simps) lemma linepath_translate_right: "linepath (a + c) (b + c) = (\x. x + c) \ linepath a b" by (auto simp: fun_eq_iff linepath_def algebra_simps) lemma integrable_on_affinity: assumes "m \ 0" "f integrable_on (cbox a b)" shows "(\x. f (m *\<^sub>R x + c)) integrable_on ((\x. (1 / m) *\<^sub>R x - ((1 / m) *\<^sub>R c)) ` cbox a b)" proof - from assms obtain I where "(f has_integral I) (cbox a b)" by (auto simp: integrable_on_def) from has_integral_affinity[OF this assms(1), of c] show ?thesis by (auto simp: integrable_on_def) qed lemma has_integral_cmul_iff: assumes "c \ 0" shows "((\x. c *\<^sub>R f x) has_integral (c *\<^sub>R I)) A \ (f has_integral I) A" using assms has_integral_cmul[of f I A c] has_integral_cmul[of "\x. c *\<^sub>R f x" "c *\<^sub>R I" A "inverse c"] by (auto simp: field_simps) lemma has_integral_affinity': fixes a :: "'a::euclidean_space" assumes "(f has_integral i) (cbox a b)" and "m > 0" shows "((\x. f(m *\<^sub>R x + c)) has_integral (i /\<^sub>R m ^ DIM('a))) (cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m))" proof (cases "cbox a b = {}") case True hence "(cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)) = {}" using \m > 0\ unfolding box_eq_empty by (auto simp: algebra_simps) with True and assms show ?thesis by simp next case False have "((\x. f (m *\<^sub>R x + c)) has_integral (1 / \m\ ^ DIM('a)) *\<^sub>R i) ((\x. (1 / m) *\<^sub>R x + - ((1 / m) *\<^sub>R c)) ` cbox a b)" using assms by (intro has_integral_affinity) auto also have "((\x. (1 / m) *\<^sub>R x + - ((1 / m) *\<^sub>R c)) ` cbox a b) = ((\x. - ((1 / m) *\<^sub>R c) + x) ` (\x. (1 / m) *\<^sub>R x) ` cbox a b)" by (simp add: image_image algebra_simps) also have "(\x. (1 / m) *\<^sub>R x) ` cbox a b = cbox ((1 / m) *\<^sub>R a) ((1 / m) *\<^sub>R b)" using \m > 0\ False by (subst image_smult_cbox) simp_all also have "(\x. - ((1 / m) *\<^sub>R c) + x) ` \ = cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)" by (subst cbox_translation [symmetric]) (simp add: field_simps vector_add_divide_simps) finally show ?thesis using \m > 0\ by (simp add: field_simps) qed lemma has_integral_affinity_iff: fixes f :: "'a :: euclidean_space \ 'b :: real_normed_vector" assumes "m > 0" shows "((\x. f (m *\<^sub>R x + c)) has_integral (I /\<^sub>R m ^ DIM('a))) (cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)) \ (f has_integral I) (cbox a b)" (is "?lhs = ?rhs") proof assume ?lhs from has_integral_affinity'[OF this, of "1 / m" "-c /\<^sub>R m"] and \m > 0\ - show ?rhs by (simp add: field_simps vector_add_divide_simps) + show ?rhs by (simp add: vector_add_divide_simps) (simp add: field_simps) next assume ?rhs from has_integral_affinity'[OF this, of m c] and \m > 0\ show ?lhs by simp qed lemma has_contour_integral_linepath_Reals_iff: fixes a b :: complex and f :: "complex \ complex" assumes "a \ Reals" "b \ Reals" "Re a < Re b" shows "(f has_contour_integral I) (linepath a b) \ ((\x. f (of_real x)) has_integral I) {Re a..Re b}" proof - from assms have [simp]: "of_real (Re a) = a" "of_real (Re b) = b" by (simp_all add: complex_eq_iff) from assms have "a \ b" by auto have "((\x. f (of_real x)) has_integral I) (cbox (Re a) (Re b)) \ ((\x. f (a + b * of_real x - a * of_real x)) has_integral I /\<^sub>R (Re b - Re a)) {0..1}" by (subst has_integral_affinity_iff [of "Re b - Re a" _ "Re a", symmetric]) (insert assms, simp_all add: field_simps scaleR_conv_of_real) also have "(\x. f (a + b * of_real x - a * of_real x)) = (\x. (f (a + b * of_real x - a * of_real x) * (b - a)) /\<^sub>R (Re b - Re a))" using \a \ b\ by (auto simp: field_simps fun_eq_iff scaleR_conv_of_real) also have "(\ has_integral I /\<^sub>R (Re b - Re a)) {0..1} \ ((\x. f (linepath a b x) * (b - a)) has_integral I) {0..1}" using assms by (subst has_integral_cmul_iff) (auto simp: linepath_def scaleR_conv_of_real algebra_simps) also have "\ \ (f has_contour_integral I) (linepath a b)" unfolding has_contour_integral_def by (intro has_integral_cong) (simp add: vector_derivative_linepath_within) finally show ?thesis by simp qed lemma contour_integrable_linepath_Reals_iff: fixes a b :: complex and f :: "complex \ complex" assumes "a \ Reals" "b \ Reals" "Re a < Re b" shows "(f contour_integrable_on linepath a b) \ (\x. f (of_real x)) integrable_on {Re a..Re b}" using has_contour_integral_linepath_Reals_iff[OF assms, of f] by (auto simp: contour_integrable_on_def integrable_on_def) lemma contour_integral_linepath_Reals_eq: fixes a b :: complex and f :: "complex \ complex" assumes "a \ Reals" "b \ Reals" "Re a < Re b" shows "contour_integral (linepath a b) f = integral {Re a..Re b} (\x. f (of_real x))" proof (cases "f contour_integrable_on linepath a b") case True thus ?thesis using has_contour_integral_linepath_Reals_iff[OF assms, of f] using has_contour_integral_integral has_contour_integral_unique by blast next case False thus ?thesis using contour_integrable_linepath_Reals_iff[OF assms, of f] by (simp add: not_integrable_contour_integral not_integrable_integral) qed lemma has_contour_integral_linepath_same_Im_iff: fixes a b :: complex and f :: "complex \ complex" assumes "Im a = Im b" "Re a < Re b" shows "(f has_contour_integral I) (linepath a b) \ ((\x. f (of_real x + Im a * \)) has_integral I) {Re a..Re b}" proof - have deriv: "vector_derivative ((\x. x - Im a * \) \ linepath a b) (at y) = b - a" for y using linepath_translate_right[of a "-Im a * \" b, symmetric] by simp have "(f has_contour_integral I) (linepath a b) \ ((\x. f (x + Im a * \)) has_contour_integral I) (linepath (a - Im a * \) (b - Im a * \))" using linepath_translate_right[of a "-Im a * \" b] deriv by (simp add: has_contour_integral) also have "\ \ ((\x. f (x + Im a * \)) has_integral I) {Re a..Re b}" using assms by (subst has_contour_integral_linepath_Reals_iff) (auto simp: complex_is_Real_iff) finally show ?thesis . qed lemma contour_integrable_linepath_same_Im_iff: fixes a b :: complex and f :: "complex \ complex" assumes "Im a = Im b" "Re a < Re b" shows "(f contour_integrable_on linepath a b) \ (\x. f (of_real x + Im a * \)) integrable_on {Re a..Re b}" using has_contour_integral_linepath_same_Im_iff[OF assms, of f] by (auto simp: contour_integrable_on_def integrable_on_def) lemma contour_integral_linepath_same_Im: fixes a b :: complex and f :: "complex \ complex" assumes "Im a = Im b" "Re a < Re b" shows "contour_integral (linepath a b) f = integral {Re a..Re b} (\x. f (x + Im a * \))" proof (cases "f contour_integrable_on linepath a b") case True thus ?thesis using has_contour_integral_linepath_same_Im_iff[OF assms, of f] using has_contour_integral_integral has_contour_integral_unique by blast next case False thus ?thesis using contour_integrable_linepath_same_Im_iff[OF assms, of f] by (simp add: not_integrable_contour_integral not_integrable_integral) qed lemmas [simp del] = div_mult_self3 div_mult_self4 div_mult_self2 div_mult_self1 lemma continuous_on_compact_bound: assumes "compact A" "continuous_on A f" obtains B where "B \ 0" "\x. x \ A \ norm (f x) \ B" proof - from assms(2,1) have "compact (f ` A)" by (rule compact_continuous_image) then obtain B where "\x\A. norm (f x) \ B" by (auto dest!: compact_imp_bounded simp: bounded_iff) hence "max B 0 \ 0" and "\x\A. norm (f x) \ max B 0" by auto thus ?thesis using that by blast qed interpretation cis: periodic_fun_simple cis "2 * pi" by standard (simp_all add: complex_eq_iff) lemma open_contains_cbox: fixes x :: "'a :: euclidean_space" assumes "open A" "x \ A" obtains a b where "cbox a b \ A" "x \ box a b" "\i\Basis. a \ i < b \ i" proof - from assms obtain R where R: "R > 0" "ball x R \ A" by (auto simp: open_contains_ball) define r :: real where "r = R / (2 * sqrt DIM('a))" from \R > 0\ have [simp]: "r > 0" by (auto simp: r_def) define d :: 'a where "d = r *\<^sub>R Topology_Euclidean_Space.One" have "cbox (x - d) (x + d) \ A" proof safe fix y assume y: "y \ cbox (x - d) (x + d)" have "dist x y = sqrt (\i\Basis. (dist (x \ i) (y \ i))\<^sup>2)" by (subst euclidean_dist_l2) (auto simp: L2_set_def) also from y have "sqrt (\i\Basis. (dist (x \ i) (y \ i))\<^sup>2) \ sqrt (\i\(Basis::'a set). r\<^sup>2)" by (intro real_sqrt_le_mono sum_mono power_mono) (auto simp: dist_norm d_def cbox_def algebra_simps) also have "\ = sqrt (DIM('a) * r\<^sup>2)" by simp also have "DIM('a) * r\<^sup>2 = (R / 2) ^ 2" by (simp add: r_def power_divide) also have "sqrt \ = R / 2" using \R > 0\ by simp also from \R > 0\ have "\ < R" by simp finally have "y \ ball x R" by simp with R show "y \ A" by blast qed thus ?thesis using that[of "x - d" "x + d"] by (auto simp: algebra_simps d_def box_def) qed lemma open_contains_box: fixes x :: "'a :: euclidean_space" assumes "open A" "x \ A" obtains a b where "box a b \ A" "x \ box a b" "\i\Basis. a \ i < b \ i" proof - from open_contains_cbox[OF assms] guess a b . with that[of a b] box_subset_cbox[of a b] show ?thesis by auto qed lemma analytic_onE_box: assumes "f analytic_on A" "s \ A" obtains a b where "Re a < Re b" "Im a < Im b" "s \ box a b" "f analytic_on box a b" proof - from assms obtain r where r: "r > 0" "f holomorphic_on ball s r" by (auto simp: analytic_on_def) with open_contains_box[of "ball s r" s] obtain a b where "box a b \ ball s r" "s \ box a b" "\i\Basis. a \ i < b \ i" by auto moreover from r have "f analytic_on ball s r" by (simp add: analytic_on_open) ultimately show ?thesis using that[of a b] analytic_on_subset[of _ "ball s r" "box a b"] by (auto simp: Basis_complex_def) qed lemma inner_image_box: assumes "(i :: 'a :: euclidean_space) \ Basis" assumes "\i\Basis. a \ i < b \ i" shows "(\x. x \ i) ` box a b = {a \ i<.. i}" proof safe fix x assume x: "x \ {a \ i<.. i}" let ?y = "(\j\Basis. (if i = j then x else (a + b) \ j / 2) *\<^sub>R j)" from x assms have "?y \ i \ (\x. x \ i) ` box a b" by (intro imageI) (auto simp: box_def algebra_simps) also have "?y \ i = (\j\Basis. (if i = j then x else (a + b) \ j / 2) * (j \ i))" by (simp add: inner_sum_left) also have "\ = (\j\Basis. if i = j then x else 0)" by (intro sum.cong) (auto simp: inner_not_same_Basis assms) also have "\ = x" using assms by simp finally show "x \ (\x. x \ i) ` box a b" . qed (insert assms, auto simp: box_def) lemma Re_image_box: assumes "Re a < Re b" "Im a < Im b" shows "Re ` box a b = {Re a<..::complex" a b] assms by (auto simp: Basis_complex_def) lemma inner_image_cbox: assumes "(i :: 'a :: euclidean_space) \ Basis" assumes "\i\Basis. a \ i \ b \ i" shows "(\x. x \ i) ` cbox a b = {a \ i..b \ i}" proof safe fix x assume x: "x \ {a \ i..b \ i}" let ?y = "(\j\Basis. (if i = j then x else a \ j) *\<^sub>R j)" from x assms have "?y \ i \ (\x. x \ i) ` cbox a b" by (intro imageI) (auto simp: cbox_def) also have "?y \ i = (\j\Basis. (if i = j then x else a \ j) * (j \ i))" by (simp add: inner_sum_left) also have "\ = (\j\Basis. if i = j then x else 0)" by (intro sum.cong) (auto simp: inner_not_same_Basis assms) also have "\ = x" using assms by simp finally show "x \ (\x. x \ i) ` cbox a b" . qed (insert assms, auto simp: cbox_def) lemma Re_image_cbox: assumes "Re a \ Re b" "Im a \ Im b" shows "Re ` cbox a b = {Re a..Re b}" using inner_image_cbox[of "1::complex" a b] assms by (auto simp: Basis_complex_def) lemma Im_image_cbox: assumes "Re a \ Re b" "Im a \ Im b" shows "Im ` cbox a b = {Im a..Im b}" using inner_image_cbox[of "\::complex" a b] assms by (auto simp: Basis_complex_def) lemma analytic_onE_cball: assumes "f analytic_on A" "s \ A" "ub > (0::real)" obtains R where "R > 0" "R < ub" "f analytic_on cball s R" proof - from assms obtain r where "r > 0" "f holomorphic_on ball s r" by (auto simp: analytic_on_def) hence "f analytic_on ball s r" by (simp add: analytic_on_open) hence "f analytic_on cball s (min (ub / 2) (r / 2))" by (rule analytic_on_subset, subst cball_subset_ball_iff) (use \r > 0\ in auto) moreover have "min (ub / 2) (r / 2) > 0" and "min (ub / 2) (r / 2) < ub" using \r > 0\ and \ub > 0\ by (auto simp: min_def) ultimately show ?thesis using that[of "min (ub / 2) (r / 2)"] by blast qed corollary analytic_pre_zeta' [analytic_intros]: assumes "f analytic_on A" "a > 0" shows "(\x. pre_zeta a (f x)) analytic_on A" using analytic_on_compose_gen[OF assms(1) analytic_pre_zeta[of a UNIV]] assms(2) by (auto simp: o_def) corollary analytic_hurwitz_zeta' [analytic_intros]: assumes "f analytic_on A" "(\x. x \ A \ f x \ 1)" "a > 0" shows "(\x. hurwitz_zeta a (f x)) analytic_on A" using analytic_on_compose_gen[OF assms(1) analytic_hurwitz_zeta[of a "-{1}"]] assms(2,3) by (auto simp: o_def) corollary analytic_zeta' [analytic_intros]: assumes "f analytic_on A" "(\x. x \ A \ f x \ 1)" shows "(\x. zeta (f x)) analytic_on A" using analytic_on_compose_gen[OF assms(1) analytic_zeta[of "-{1}"]] assms(2) by (auto simp: o_def) lemma logderiv_zeta_analytic: "(\s. deriv zeta s / zeta s) analytic_on {s. Re s \ 1} - {1}" using zeta_Re_ge_1_nonzero by (auto intro!: analytic_intros) lemma cis_pi_half [simp]: "cis (pi / 2) = \" by (simp add: complex_eq_iff) lemma mult_real_sqrt: "x \ 0 \ x * sqrt y = sqrt (x ^ 2 * y)" by (simp add: real_sqrt_mult) lemma arcsin_pos: "x \ {0<..1} \ arcsin x > 0" using arcsin_less_arcsin[of 0 x] by simp lemmas analytic_imp_holomorphic' = holomorphic_on_subset[OF analytic_imp_holomorphic] lemma residue_simple': assumes "open s" "0 \ s" "f holomorphic_on s" shows "residue (\w. f w / w) 0 = f 0" using residue_simple[of s 0 f] assms by simp lemma fds_converges_cong: assumes "eventually (\n. fds_nth f n = fds_nth g n) at_top" "s = s'" shows "fds_converges f s \ fds_converges g s'" unfolding fds_converges_def by (intro summable_cong eventually_mono[OF assms(1)]) (simp_all add: assms) lemma fds_abs_converges_cong: assumes "eventually (\n. fds_nth f n = fds_nth g n) at_top" "s = s'" shows "fds_abs_converges f s \ fds_abs_converges g s'" unfolding fds_abs_converges_def by (intro summable_cong eventually_mono[OF assms(1)]) (simp_all add: assms) lemma conv_abscissa_cong: assumes "eventually (\n. fds_nth f n = fds_nth g n) at_top" shows "conv_abscissa f = conv_abscissa g" proof - have "fds_converges f = fds_converges g" by (intro ext fds_converges_cong assms refl) thus ?thesis by (simp add: conv_abscissa_def) qed lemma abs_conv_abscissa_cong: assumes "eventually (\n. fds_nth f n = fds_nth g n) at_top" shows "abs_conv_abscissa f = abs_conv_abscissa g" proof - have "fds_abs_converges f = fds_abs_converges g" by (intro ext fds_abs_converges_cong assms refl) thus ?thesis by (simp add: abs_conv_abscissa_def) qed definition fds_remainder where "fds_remainder m = fds_subseries (\n. n > m)" lemma fds_nth_remainder: "fds_nth (fds_remainder m f) = (\n. if n > m then fds_nth f n else 0)" by (simp add: fds_remainder_def fds_subseries_def fds_nth_fds') lemma fds_converges_remainder_iff [simp]: "fds_converges (fds_remainder m f) s \ fds_converges f s" by (intro fds_converges_cong eventually_mono[OF eventually_gt_at_top[of m]]) (auto simp: fds_nth_remainder) lemma fds_abs_converges_remainder_iff [simp]: "fds_abs_converges (fds_remainder m f) s \ fds_abs_converges f s" by (intro fds_abs_converges_cong eventually_mono[OF eventually_gt_at_top[of m]]) (auto simp: fds_nth_remainder) lemma fds_converges_remainder [intro]: "fds_converges f s \ fds_converges (fds_remainder m f) s" and fds_abs_converges_remainder [intro]: "fds_abs_converges f s \ fds_abs_converges (fds_remainder m f) s" by simp_all lemma conv_abscissa_remainder [simp]: "conv_abscissa (fds_remainder m f) = conv_abscissa f" by (intro conv_abscissa_cong eventually_mono[OF eventually_gt_at_top[of m]]) (auto simp: fds_nth_remainder) lemma abs_conv_abscissa_remainder [simp]: "abs_conv_abscissa (fds_remainder m f) = abs_conv_abscissa f" by (intro abs_conv_abscissa_cong eventually_mono[OF eventually_gt_at_top[of m]]) (auto simp: fds_nth_remainder) lemma eval_fds_remainder: "eval_fds (fds_remainder m f) s = (\n. fds_nth f (n + Suc m) / nat_power (n + Suc m) s)" (is "_ = suminf (\n. ?f (n + Suc m))") proof (cases "fds_converges f s") case False hence "\fds_converges (fds_remainder m f) s" by simp hence "(\x. (\n. fds_nth (fds_remainder m f) n / nat_power n s) sums x) = (\_. False)" by (auto simp: fds_converges_def summable_def) hence "eval_fds (fds_remainder m f) s = (THE _. False)" by (simp add: eval_fds_def suminf_def) moreover from False have "\summable (\n. ?f (n + Suc m))" unfolding fds_converges_def by (subst summable_iff_shift) auto hence "(\x. (\n. ?f (n + Suc m)) sums x) = (\_. False)" by (auto simp: summable_def) hence "suminf (\n. ?f (n + Suc m)) = (THE _. False)" by (simp add: suminf_def) ultimately show ?thesis by simp next case True hence *: "fds_converges (fds_remainder m f) s" by simp have "eval_fds (fds_remainder m f) s = (\n. fds_nth (fds_remainder m f) n / nat_power n s)" unfolding eval_fds_def .. also have "\ = (\n. fds_nth (fds_remainder m f) (n + Suc m) / nat_power (n + Suc m) s)" using * unfolding fds_converges_def by (subst suminf_minus_initial_segment) (auto simp: fds_nth_remainder) also have "(\n. fds_nth (fds_remainder m f) (n + Suc m)) = (\n. fds_nth f (n + Suc m))" by (intro ext) (auto simp: fds_nth_remainder) finally show ?thesis . qed lemma fds_truncate_plus_remainder: "fds_truncate m f + fds_remainder m f = f" by (intro fds_eqI) (auto simp: fds_truncate_def fds_remainder_def fds_subseries_def) lemma holomorphic_fds_eval' [holomorphic_intros]: assumes "g holomorphic_on A" "\x. x \ A \ Re (g x) > conv_abscissa f" shows "(\x. eval_fds f (g x)) holomorphic_on A" using holomorphic_on_compose_gen[OF assms(1) holomorphic_fds_eval[OF order.refl, of f]] assms(2) by (auto simp: o_def) lemma analytic_fds_eval' [analytic_intros]: assumes "g analytic_on A" "\x. x \ A \ Re (g x) > conv_abscissa f" shows "(\x. eval_fds f (g x)) analytic_on A" using analytic_on_compose_gen[OF assms(1) analytic_fds_eval[OF order.refl, of f]] assms(2) by (auto simp: o_def) lemma homotopic_loopsI: fixes h :: "real \ real \ _" assumes "continuous_on ({0..1} \ {0..1}) h" "h ` ({0..1} \ {0..1}) \ s" "\x. x \ {0..1} \ h (0, x) = p x" "\x. x \ {0..1} \ h (1, x) = q x" "\x. x \ {0..1} \ pathfinish (h \ Pair x) = pathstart (h \ Pair x)" shows "homotopic_loops s p q" using assms unfolding homotopic_loops by (intro exI[of _ h]) auto lemma continuous_on_linepath [continuous_intros]: assumes "continuous_on A a" "continuous_on A b" "continuous_on A f" shows "continuous_on A (\x. linepath (a x) (b x) (f x))" using assms by (auto simp: linepath_def intro!: continuous_intros assms) lemma continuous_on_part_circlepath [continuous_intros]: assumes "continuous_on A c" "continuous_on A r" "continuous_on A a" "continuous_on A b" "continuous_on A f" shows "continuous_on A (\x. part_circlepath (c x) (r x) (a x) (b x) (f x))" using assms by (auto simp: part_circlepath_def intro!: continuous_intros assms) lemma homotopic_loops_part_circlepath: assumes "sphere c r \ A" and "r \ 0" and "b1 = a1 + 2 * of_int k * pi" and "b2 = a2 + 2 * of_int k * pi" shows "homotopic_loops A (part_circlepath c r a1 b1) (part_circlepath c r a2 b2)" proof - define h where "h = (\(x,y). part_circlepath c r (linepath a1 a2 x) (linepath b1 b2 x) y)" show ?thesis proof (rule homotopic_loopsI) show "continuous_on ({0..1} \ {0..1}) h" by (auto simp: h_def case_prod_unfold intro!: continuous_intros) next from assms have "h ` ({0..1} \ {0..1}) \ sphere c r" by (auto simp: h_def part_circlepath_def dist_norm norm_mult) also have "\ \ A" by fact finally show "h ` ({0..1} \ {0..1}) \ A" . next fix x :: real assume x: "x \ {0..1}" show "h (0, x) = part_circlepath c r a1 b1 x" and "h (1, x) = part_circlepath c r a2 b2 x" by (simp_all add: h_def linepath_def) have "cis (pi * (real_of_int k * 2)) = 1" using cis.plus_of_int[of 0 k] by (simp add: algebra_simps) thus "pathfinish (h \ Pair x) = pathstart (h \ Pair x)" by (simp add: h_def o_def exp_eq_polar linepath_def algebra_simps cis_mult [symmetric] cis_divide [symmetric] assms) qed qed lemma homotopic_pathsI: fixes h :: "real \ real \ _" assumes "continuous_on ({0..1} \ {0..1}) h" assumes "h ` ({0..1} \ {0..1}) \ s" assumes "\x. x \ {0..1} \ h (0, x) = p x" assumes "\x. x \ {0..1} \ h (1, x) = q x" assumes "\x. x \ {0..1} \ pathstart (h \ Pair x) = pathstart p" assumes "\x. x \ {0..1} \ pathfinish (h \ Pair x) = pathfinish p" shows "homotopic_paths s p q" using assms unfolding homotopic_paths by (intro exI[of _ h]) auto lemma part_circlepath_conv_subpath: "part_circlepath c r a b = subpath (a / (2*pi)) (b / (2*pi)) (circlepath c r)" by (simp add: part_circlepath_def circlepath_def subpath_def linepath_def algebra_simps exp_eq_polar) lemma homotopic_paths_part_circlepath: assumes "a \ b" "b \ c" assumes "path_image (part_circlepath C r a c) \ A" "r \ 0" shows "homotopic_paths A (part_circlepath C r a c) (part_circlepath C r a b +++ part_circlepath C r b c)" (is "homotopic_paths _ ?g (?h1 +++ ?h2)") proof (cases "a = c") case False with assms have "a < c" by simp define slope where "slope = (b - a) / (c - a)" from assms and \a < c\ have slope: "slope \ {0..1}" by (auto simp: field_simps slope_def) define f :: "real \ real" where "f = linepath 0 slope +++ linepath slope 1" show ?thesis proof (rule homotopic_paths_reparametrize) fix t :: real assume t: "t \ {0..1}" show "(?h1 +++ ?h2) t = ?g (f t)" proof (cases "t \ 1 / 2") case True hence "?g (f t) = C + r * cis ((1 - f t) * a + f t * c)" by (simp add: joinpaths_def part_circlepath_def exp_eq_polar linepath_def) also from True \a < c\ have "(1 - f t) * a + f t * c = (1 - 2 * t) * a + 2 * t * b" unfolding f_def slope_def linepath_def joinpaths_def by (simp add: divide_simps del: div_mult_self3 div_mult_self4 div_mult_self2 div_mult_self1) (simp add: algebra_simps)? also from True have "C + r * cis \ = (?h1 +++ ?h2) t" by (simp add: joinpaths_def part_circlepath_def exp_eq_polar linepath_def) finally show ?thesis .. next case False hence "?g (f t) = C + r * cis ((1 - f t) * a + f t * c)" by (simp add: joinpaths_def part_circlepath_def exp_eq_polar linepath_def) also from False \a < c\ have "(1 - f t) * a + f t * c = (2 - 2 * t) * b + (2 * t - 1) * c" unfolding f_def slope_def linepath_def joinpaths_def by (simp add: divide_simps del: div_mult_self3 div_mult_self4 div_mult_self2 div_mult_self1) (simp add: algebra_simps)? also from False have "C + r * cis \ = (?h1 +++ ?h2) t" by (simp add: joinpaths_def part_circlepath_def exp_eq_polar linepath_def) finally show ?thesis .. qed next from slope have "path_image f \ {0..1}" by (auto simp: f_def path_image_join closed_segment_eq_real_ivl) thus "f ` {0..1} \ {0..1}" by (simp add: path_image_def) next have "path f" unfolding f_def by auto thus "continuous_on {0..1} f" by (simp add: path_def) qed (insert assms, auto simp: f_def joinpaths_def linepath_def) next case [simp]: True with assms have [simp]: "b = c" by auto have "part_circlepath C r c c +++ part_circlepath C r c c = part_circlepath C r c c" by (simp add: fun_eq_iff joinpaths_def part_circlepath_def) thus ?thesis using assms by simp qed lemma has_contour_integral_mirror_iff: assumes "valid_path g" shows "(f has_contour_integral I) (-g) \ ((\x. -f (- x)) has_contour_integral I) g" proof - from assms have "g piecewise_differentiable_on {0..1}" by (auto simp: valid_path_def piecewise_C1_imp_differentiable) then obtain S where S: "finite S" "\x. x \ {0..1} - S \ g differentiable at x within {0..1}" unfolding piecewise_differentiable_on_def by blast have S': "g differentiable at x" if "x \ {0..1} - ({0, 1} \ S)" for x proof - from that have "x \ interior {0..1}" by auto with S(2)[of x] that show ?thesis by (auto simp: at_within_interior[of _ "{0..1}"]) qed have "(f has_contour_integral I) (-g) \ ((\x. f (- g x) * vector_derivative (-g) (at x)) has_integral I) {0..1}" by (simp add: has_contour_integral) also have "\ \ ((\x. -f (- g x) * vector_derivative g (at x)) has_integral I) {0..1}" by (intro has_integral_spike_finite_eq[of "S \ {0, 1}"]) (insert \finite S\ S', auto simp: o_def fun_Compl_def) also have "\ \ ((\x. -f (-x)) has_contour_integral I) g" by (simp add: has_contour_integral) finally show ?thesis . qed lemma contour_integral_on_mirror_iff: assumes "valid_path g" shows "f contour_integrable_on (-g) \ (\x. -f (- x)) contour_integrable_on g" by (auto simp: contour_integrable_on_def has_contour_integral_mirror_iff assms) lemma contour_integral_mirror: assumes "valid_path g" shows "contour_integral (-g) f = contour_integral g (\x. -f (- x))" proof (cases "f contour_integrable_on (-g)") case True then obtain I where I: "(f has_contour_integral I) (-g)" by (auto simp: contour_integrable_on_def) also note has_contour_integral_mirror_iff[OF assms] finally have "((\x. - f (- x)) has_contour_integral I) g" . with I show ?thesis using contour_integral_unique by blast next case False hence "\(\x. -f (-x)) contour_integrable_on g" by (auto simp: contour_integral_on_mirror_iff assms) from False and this show ?thesis by (simp add: not_integrable_contour_integral) qed lemma contour_integrable_neg_iff: "(\x. -f x) contour_integrable_on g \ f contour_integrable_on g" using contour_integrable_neg[of f g] contour_integrable_neg[of "\x. -f x" g] by auto lemma contour_integral_neg: shows "contour_integral g (\x. -f x) = -contour_integral g f" proof (cases "f contour_integrable_on g") case True thus ?thesis by (simp add: contour_integral_neg) next case False hence "\(\x. -f x) contour_integrable_on g" by (simp add: contour_integrable_neg_iff) with False show ?thesis by (simp add: not_integrable_contour_integral) qed lemma minus_cis: "-cis x = cis (x + pi)" by (simp add: complex_eq_iff) lemma path_image_part_circlepath_subset: assumes "a \ a'" "a' \ b'" "b' \ b" shows "path_image (part_circlepath c r a' b') \ path_image (part_circlepath c r a b)" using assms by (subst (1 2) path_image_part_circlepath) auto lemma part_circlepath_mirror: assumes "a' = a + pi + 2 * pi * of_int k" "b' = b + pi + 2 * pi * of_int k" "c' = -c" shows "-part_circlepath c r a b = part_circlepath c' r a' b'" proof fix x :: real have "part_circlepath c' r a' b' x = c' + r * cis (linepath a b x + pi + k * (2 * pi))" by (simp add: part_circlepath_def exp_eq_polar assms linepath_translate_right mult_ac) also have "cis (linepath a b x + pi + k * (2 * pi)) = cis (linepath a b x + pi)" by (rule cis.plus_of_int) also have "\ = -cis (linepath a b x)" by (simp add: minus_cis) also have "c' + r * \ = -part_circlepath c r a b x" by (simp add: part_circlepath_def assms exp_eq_polar) finally show "(- part_circlepath c r a b) x = part_circlepath c' r a' b' x" by simp qed lemma path_mirror [intro]: "path (g :: _ \ 'b::topological_group_add) \ path (-g)" by (auto simp: path_def intro!: continuous_intros) lemma path_mirror_iff [simp]: "path (-g :: _ \ 'b::topological_group_add) \ path g" using path_mirror[of g] path_mirror[of "-g"] by (auto simp: fun_Compl_def) lemma valid_path_mirror [intro]: "valid_path g \ valid_path (-g)" by (auto simp: valid_path_def fun_Compl_def piecewise_C1_differentiable_neg) lemma valid_path_mirror_iff [simp]: "valid_path (-g) \ valid_path g" using valid_path_mirror[of g] valid_path_mirror[of "-g"] by (auto simp: fun_Compl_def) lemma pathstart_mirror [simp]: "pathstart (-g) = -pathstart g" and pathfinish_mirror [simp]: "pathfinish (-g) = -pathfinish g" by (simp_all add: pathstart_def pathfinish_def) lemma path_image_mirror: "path_image (-g) = uminus ` path_image g" by (auto simp: path_image_def) lemma contour_integral_bound_part_circlepath: assumes "f contour_integrable_on part_circlepath c r a b" assumes "B \ 0" "r \ 0" "\x. x \ path_image (part_circlepath c r a b) \ norm (f x) \ B" shows "norm (contour_integral (part_circlepath c r a b) f) \ B * r * \b - a\" proof - let ?I = "integral {0..1} (\x. f (part_circlepath c r a b x) * \ * of_real (r * (b - a)) * exp (\ * linepath a b x))" have "norm ?I \ integral {0..1} (\x::real. B * 1 * (r * \b - a\) * 1)" proof (rule integral_norm_bound_integral, goal_cases) case 1 with assms(1) show ?case by (simp add: contour_integrable_on vector_derivative_part_circlepath mult_ac) next case (3 x) with assms(2-) show ?case unfolding norm_mult norm_of_real abs_mult by (intro mult_mono) (auto simp: path_image_def) qed auto also have "?I = contour_integral (part_circlepath c r a b) f" by (simp add: contour_integral_integral vector_derivative_part_circlepath mult_ac) finally show ?thesis by simp qed lemma contour_integral_spike_finite_simple_path: assumes "finite A" "simple_path g" "g = g'" "\x. x \ path_image g - A \ f x = f' x" shows "contour_integral g f = contour_integral g' f'" unfolding contour_integral_integral proof (rule integral_spike) have "finite (g -` A \ {0<..<1})" using \simple_path g\ \finite A\ by (intro finite_vimage_IntI simple_path_inj_on) auto hence "finite ({0, 1} \ g -` A \ {0<..<1})" by auto thus "negligible ({0, 1} \ g -` A \ {0<..<1})" by (rule negligible_finite) next fix x assume "x \ {0..1} - ({0, 1} \ g -` A \ {0<..<1})" hence "g x \ path_image g - A" by (auto simp: path_image_def) from assms(4)[OF this] and assms(3) show "f' (g' x) * vector_derivative g' (at x) = f (g x) * vector_derivative g (at x)" by simp qed proposition contour_integral_bound_part_circlepath_strong: assumes fi: "f contour_integrable_on part_circlepath z r s t" and "finite k" and le: "0 \ B" "0 < r" "s \ t" and B: "\x. x \ path_image(part_circlepath z r s t) - k \ norm(f x) \ B" shows "cmod (contour_integral (part_circlepath z r s t) f) \ B * r * (t - s)" proof - from fi have "(f has_contour_integral contour_integral (part_circlepath z r s t) f) (part_circlepath z r s t)" by (rule has_contour_integral_integral) from has_contour_integral_bound_part_circlepath_strong[OF this assms(2-)] show ?thesis by auto qed lemma cos_le_zero: assumes "x \ {pi/2..3*pi/2}" shows "cos x \ 0" proof - have "cos x = -cos (x - pi)" by (simp add: cos_diff) moreover from assms have "cos (x - pi) \ 0" by (intro cos_ge_zero) auto ultimately show ?thesis by simp qed lemma cos_le_zero': "x \ {-3*pi/2..-pi/2} \ cos x \ 0" using cos_le_zero[of "-x"] by simp lemma cis_minus_pi_half [simp]: "cis (- (pi / 2)) = -\" by (simp add: complex_eq_iff) lemma winding_number_join_pos_combined': "\valid_path \1 \ z \ path_image \1 \ 0 < Re (winding_number \1 z); valid_path \2 \ z \ path_image \2 \ 0 < Re (winding_number \2 z); pathfinish \1 = pathstart \2\ \ valid_path(\1 +++ \2) \ z \ path_image(\1 +++ \2) \ 0 < Re(winding_number(\1 +++ \2) z)" by (simp add: valid_path_join path_image_join winding_number_join valid_path_imp_path) lemma Union_atLeastAtMost_real_of_nat: assumes "a < b" shows "(\n\{a.. {real a..real b}" thus "x \ (\n\{a.. real a" "x < real b" by simp_all hence "x \ real (nat \x\)" "x \ real (Suc (nat \x\))" by linarith+ moreover from x have "nat \x\ \ a" "nat \x\ < b" by linarith+ ultimately have "\n\{a.. {real n..real (n + 1)}" by (intro bexI[of _ "nat \x\"]) simp_all thus ?thesis by blast qed qed auto lemma nat_sum_has_integral_floor: fixes f :: "nat \ 'a :: banach" assumes mn: "m < n" shows "((\x. f (nat \x\)) has_integral sum f {m..i. {real i..real (Suc i)}) ` {m..x. f (nat \x\)) has_integral (\X\D. f (nat \Inf X\))) {real m..real n}" proof (rule has_integral_combine_division) fix X assume X: "X \ D" have "nat \x\ = nat \Inf X\" if "x \ X - {Sup X}" for x using that X by (auto simp: D_def nat_eq_iff floor_eq_iff) hence "((\x. f (nat \x\)) has_integral f (nat \Inf X\)) X \ ((\x. f (nat \Inf X\)) has_integral f (nat \Inf X\)) X" using X by (intro has_integral_spike_eq[of "{Sup X}"]) auto also from X have "\" using has_integral_const_real[of "f (nat \Inf X\)" "Inf X" "Sup X"] by (auto simp: D_def) finally show "((\x. f (nat \x\)) has_integral f (nat \Inf X\)) X" . qed fact+ also have "(\X\D. f (nat \Inf X\)) = (\k\{m.. 'a :: banach" assumes mn: "m < n" shows "((\x. f (nat \x\)) has_integral sum f {m<..n}) {real m..real n}" proof - define D where "D = (\i. {real i..real (Suc i)}) ` {m..x. f (nat \x\)) has_integral (\X\D. f (nat \Sup X\))) {real m..real n}" proof (rule has_integral_combine_division) fix X assume X: "X \ D" have "nat \x\ = nat \Sup X\" if "x \ X - {Inf X}" for x using that X by (auto simp: D_def nat_eq_iff ceiling_eq_iff) hence "((\x. f (nat \x\)) has_integral f (nat \Sup X\)) X \ ((\x. f (nat \Sup X\)) has_integral f (nat \Sup X\)) X" using X by (intro has_integral_spike_eq[of "{Inf X}"]) auto also from X have "\" using has_integral_const_real[of "f (nat \Sup X\)" "Inf X" "Sup X"] by (auto simp: D_def) finally show "((\x. f (nat \x\)) has_integral f (nat \Sup X\)) X" . qed fact+ also have "(\X\D. f (nat \Sup X\)) = (\k\{m.. = (\k\{m<..n}. f k)" by (intro sum.reindex_bij_witness[of _ "\x. x - 1" Suc]) auto finally show ?thesis . qed lemma zeta_partial_sum_le: fixes x :: real and m :: nat assumes x: "x \ {0<..1}" shows "(\k=1..m. real k powr (x - 1)) \ real m powr x / x" proof - consider "m = 0" | "m = 1" | "m > 1" by force thus ?thesis proof cases assume m: "m > 1" hence "{1..m} = insert 1 {1<..m}" by auto also have "(\k\\. real k powr (x - 1)) = 1 + (\k\{1<..m}. real k powr (x - 1))" by simp also have "(\k\{1<..m}. real k powr (x - 1)) \ real m powr x / x - 1 / x" proof (rule has_integral_le) show "((\t. (nat \t\) powr (x - 1)) has_integral (\n\{1<..m}. n powr (x - 1))) {real 1..m}" using m by (intro nat_sum_has_integral_ceiling) auto next have "((\t. t powr (x - 1)) has_integral (real m powr x / x - real 1 powr x / x)) {real 1..real m}" by (intro fundamental_theorem_of_calculus) (insert x m, auto simp flip: has_field_derivative_iff_has_vector_derivative intro!: derivative_eq_intros) thus "((\t. t powr (x - 1)) has_integral (real m powr x / x - 1 / x)) {real 1..real m}" by simp qed (insert x, auto intro!: powr_mono2') also have "1 + (real m powr x / x - 1 / x) \ real m powr x / x" using x by (simp add: field_simps) finally show ?thesis by simp qed (use assms in auto) qed lemma zeta_partial_sum_le': fixes x :: real and m :: nat assumes x: "x > 0" and m: "m > 0" shows "(\n=1..m. real n powr (x - 1)) \ m powr x * (1 / x + 1 / m)" proof (cases "x > 1") case False with assms have "(\n=1..m. real n powr (x - 1)) \ m powr x / x" by (intro zeta_partial_sum_le) auto also have "\ \ m powr x * (1 / x + 1 / m)" using assms by (simp add: field_simps) finally show ?thesis . next case True have "(\n\{1..m}. n powr (x - 1)) = (\n\insert m {0.. = m powr (x - 1) + (\n\{0..n\{0.. real m powr x / x" proof (rule has_integral_le) show "((\t. (nat \t\) powr (x - 1)) has_integral (\n\{0..t. t powr (x - 1)) has_integral (real m powr x / x)) {real 0..real m}" using has_integral_powr_from_0[of "x - 1"] x by auto next fix t assume "t \ {real 0..real m}" with \x > 1\ show "real (nat \t\) powr (x - 1) \ t powr (x - 1)" by (cases "t = 0") (auto intro: powr_mono2) qed also have "m powr (x - 1) + m powr x / x = m powr x * (1 / x + 1 / m)" using m x by (simp add: powr_diff field_simps) finally show ?thesis by simp qed lemma natfun_bigo_1E: assumes "(f :: nat \ _) \ O(\_. 1)" obtains C where "C \ lb" "\n. norm (f n) \ C" proof - from assms obtain C N where "\n\N. norm (f n) \ C" by (auto elim!: landau_o.bigE simp: eventually_at_top_linorder) hence *: "norm (f n) \ Max ({C, lb} \ (norm ` f ` {.. N") (subst Max_ge_iff; force simp: image_iff)+ moreover have "Max ({C, lb} \ (norm ` f ` {.. lb" by (intro Max.coboundedI) auto ultimately show ?thesis using that by blast qed lemma natfun_bigo_iff_Bseq: "f \ O(\_. 1) \ Bseq f" proof assume "Bseq f" then obtain C where "C > 0" "\n. norm (f n) \ C" by (auto simp: Bseq_def) thus "f \ O(\_. 1)" by (intro bigoI[of _ C]) auto next assume "f \ O(\_. 1)" from natfun_bigo_1E[OF this, where lb = 1] obtain C where "C \ 1" "\n. norm (f n) \ C" by auto thus "Bseq f" by (auto simp: Bseq_def intro!: exI[of _ C]) qed lemma enn_decreasing_sum_le_set_nn_integral: fixes f :: "real \ ennreal" assumes decreasing: "\x y. 0 \ x \ x \ y \ f y \ f x" shows "(\n. f (real (Suc n))) \ set_nn_integral lborel {0..} f" proof - have "(\n. (f (Suc n))) = (\n. \\<^sup>+x\{real n<..real (Suc n)}. (f (Suc n)) \lborel)" by (subst nn_integral_cmult_indicator) auto also have "nat \x\ = Suc n" if "x \ {real n<..real (Suc n)}" for x n using that by (auto simp: nat_eq_iff ceiling_eq_iff) hence "(\n. \\<^sup>+x\{real n<..real (Suc n)}. (f (Suc n)) \lborel) = (\n. \\<^sup>+x\{real n<..real (Suc n)}. (f (real (nat \x\))) \lborel)" by (intro suminf_cong nn_integral_cong) (auto simp: indicator_def) also have "\ = (\\<^sup>+x\(\i. {real i<..real (Suc i)}). (f (nat \x::real\)) \lborel)" by (subst nn_integral_disjoint_family) (auto simp: disjoint_family_on_def) also have "\ \ (\\<^sup>+x\{0..}. (f x) \lborel)" by (intro nn_integral_mono) (auto simp: indicator_def intro!: decreasing) finally show ?thesis . qed (* TODO replace version in library *) lemma nn_integral_has_integral_lebesgue: fixes f :: "'a::euclidean_space \ real" assumes nonneg: "\x. x \ \ \ 0 \ f x" and I: "(f has_integral I) \" shows "integral\<^sup>N lborel (\x. indicator \ x * f x) = I" proof - from I have "(\x. indicator \ x *\<^sub>R f x) \ lebesgue \\<^sub>M borel" by (rule has_integral_implies_lebesgue_measurable) then obtain f' :: "'a \ real" where [measurable]: "f' \ borel \\<^sub>M borel" and eq: "AE x in lborel. indicator \ x * f x = f' x" by (auto dest: completion_ex_borel_measurable_real) from I have "((\x. abs (indicator \ x * f x)) has_integral I) UNIV" using nonneg by (simp add: indicator_def if_distrib[of "\x. x * f y" for y] cong: if_cong) also have "((\x. abs (indicator \ x * f x)) has_integral I) UNIV \ ((\x. abs (f' x)) has_integral I) UNIV" using eq by (intro has_integral_AE) auto finally have "integral\<^sup>N lborel (\x. abs (f' x)) = I" by (rule nn_integral_has_integral_lborel[rotated 2]) auto also have "integral\<^sup>N lborel (\x. abs (f' x)) = integral\<^sup>N lborel (\x. abs (indicator \ x * f x))" using eq by (intro nn_integral_cong_AE) auto also have "(\x. abs (indicator \ x * f x)) = (\x. indicator \ x * f x)" using nonneg by (auto simp: indicator_def fun_eq_iff) finally show ?thesis . qed lemma decreasing_sum_le_integral: fixes f :: "real \ real" assumes nonneg: "\x. x \ 0 \ f x \ 0" assumes decreasing: "\x y. 0 \ x \ x \ y \ f y \ f x" assumes integral: "(f has_integral I) {0..}" shows "summable (\i. f (real (Suc i)))" and "suminf (\i. f (real (Suc i))) \ I" proof - have [simp]: "I \ 0" by (intro has_integral_nonneg[OF integral] nonneg) auto have "(\n. ennreal (f (Suc n))) = (\n. \\<^sup>+x\{real n<..real (Suc n)}. ennreal (f (Suc n)) \lborel)" by (subst nn_integral_cmult_indicator) auto also have "nat \x\ = Suc n" if "x \ {real n<..real (Suc n)}" for x n using that by (auto simp: nat_eq_iff ceiling_eq_iff) hence "(\n. \\<^sup>+x\{real n<..real (Suc n)}. ennreal (f (Suc n)) \lborel) = (\n. \\<^sup>+x\{real n<..real (Suc n)}. ennreal (f (real (nat \x\))) \lborel)" by (intro suminf_cong nn_integral_cong) (auto simp: indicator_def) also have "\ = (\\<^sup>+x\(\i. {real i<..real (Suc i)}). ennreal (f (nat \x::real\)) \lborel)" by (subst nn_integral_disjoint_family) (auto simp: disjoint_family_on_def intro!: measurable_completion) also have "\ \ (\\<^sup>+x\{0..}. ennreal (f x) \lborel)" by (intro nn_integral_mono) (auto simp: indicator_def nonneg intro!: decreasing) also have "\ = (\\<^sup>+ x. ennreal (indicat_real {0..} x * f x) \lborel)" by (intro nn_integral_cong) (auto simp: indicator_def) also have "\ = ennreal I" using nn_integral_has_integral_lebesgue[OF nonneg integral] by (auto simp: nonneg) finally have *: "(\n. ennreal (f (Suc n))) \ ennreal I" . from * show summable: "summable (\i. f (real (Suc i)))" by (intro summable_suminf_not_top) (auto simp: top_unique intro: nonneg) note * also from summable have "(\n. ennreal (f (Suc n))) = ennreal (\n. f (Suc n))" by (subst suminf_ennreal2) (auto simp: o_def nonneg) finally show "(\n. f (real (Suc n))) \ I" by (subst (asm) ennreal_le_iff) auto qed lemma decreasing_sum_le_integral': fixes f :: "real \ real" assumes "\x. x \ 0 \ f x \ 0" assumes "\x y. 0 \ x \ x \ y \ f y \ f x" assumes "(f has_integral I) {0..}" shows "summable (\i. f (real i))" and "suminf (\i. f (real i)) \ f 0 + I" proof - have "summable ((\i. f (real (Suc i))))" using decreasing_sum_le_integral[OF assms] by (simp add: o_def) thus *: "summable (\i. f (real i))" by (subst (asm) summable_Suc_iff) have "(\n. f (real (Suc n))) \ I" by (intro decreasing_sum_le_integral assms) thus "suminf (\i. f (real i)) \ f 0 + I" using * by (subst (asm) suminf_split_head) auto qed lemma norm_suminf_le: assumes "\n. norm (f n :: 'a :: banach) \ g n" "summable g" shows "norm (suminf f) \ suminf g" proof - have *: "summable (\n. norm (f n))" using assms by (intro summable_norm summable_comparison_test[OF _ assms(2)] exI[of _ 0]) auto hence "norm (suminf f) \ (\n. norm (f n))" by (intro summable_norm) auto also have "\ \ suminf g" by (intro suminf_le * assms allI) finally show ?thesis . qed lemma of_nat_powr_neq_1_complex [simp]: assumes "n > 1" "Re s \ 0" shows "of_nat n powr s \ (1::complex)" proof - have "norm (of_nat n powr s) = real n powr Re s" by (simp add: norm_powr_real_powr) also have "\ \ 1" using assms by (auto simp: powr_def) finally show ?thesis by auto qed lemma abs_summable_on_uminus_iff: "(\x. -f x) abs_summable_on A \ f abs_summable_on A" using abs_summable_on_uminus[of f A] abs_summable_on_uminus[of "\x. -f x" A] by auto lemma abs_summable_on_cmult_right_iff: fixes f :: "'a \ 'b :: {banach, real_normed_field, second_countable_topology}" assumes "c \ 0" shows "(\x. c * f x) abs_summable_on A \ f abs_summable_on A" using assms abs_summable_on_cmult_right[of c f A] abs_summable_on_cmult_right[of "inverse c" "\x. c * f x" A] by (auto simp: field_simps) lemma abs_summable_on_cmult_left_iff: fixes f :: "'a \ 'b :: {banach, real_normed_field, second_countable_topology}" assumes "c \ 0" shows "(\x. f x * c) abs_summable_on A \ f abs_summable_on A" using assms abs_summable_on_cmult_left[of c f A] abs_summable_on_cmult_left[of "inverse c" "\x. f x * c" A] by (auto simp: field_simps) lemma fds_logderiv_completely_multiplicative: fixes f :: "'a :: {real_normed_field} fds" assumes "completely_multiplicative_function (fds_nth f)" "fds_nth f 1 \ 0" shows "fds_deriv f / f = - fds (\n. fds_nth f n * mangoldt n)" proof - have "fds_deriv f / f = - fds (\n. fds_nth f n * mangoldt n) * f / f" using completely_multiplicative_fds_deriv[of "fds_nth f"] assms by simp also have "\ = - fds (\n. fds_nth f n * mangoldt n)" using assms by (simp add: divide_fds_def fds_right_inverse) finally show ?thesis . qed lemma fds_nth_logderiv_completely_multiplicative: fixes f :: "'a :: {real_normed_field} fds" assumes "completely_multiplicative_function (fds_nth f)" "fds_nth f 1 \ 0" shows "fds_nth (fds_deriv f / f) n = -fds_nth f n * mangoldt n" using assms by (subst fds_logderiv_completely_multiplicative) (simp_all add: fds_nth_fds') lemma eval_fds_logderiv_completely_multiplicative: fixes s :: "'a :: dirichlet_series" and l :: 'a and f :: "'a fds" defines "h \ fds_deriv f / f" assumes "completely_multiplicative_function (fds_nth f)" and [simp]: "fds_nth f 1 \ 0" assumes "s \ 1 > abs_conv_abscissa f" shows "(\p. of_real (ln (real p)) * (1 / (1 - fds_nth f p / nat_power p s) - 1)) abs_summable_on {p. prime p}" (is ?th1) and "eval_fds h s = -(\\<^sub>ap | prime p. of_real (ln (real p)) * (1 / (1 - fds_nth f p / nat_power p s) - 1))" (is ?th2) proof - let ?P = "{p::nat. prime p}" interpret f: completely_multiplicative_function "fds_nth f" by fact have "fds_abs_converges h s" using abs_conv_abscissa_completely_multiplicative_log_deriv[OF assms(2)] assms by (intro fds_abs_converges) auto hence *: "(\n. fds_nth h n / nat_power n s) abs_summable_on UNIV" by (auto simp: h_def fds_abs_converges_altdef') note * also have "(\n. fds_nth h n / nat_power n s) abs_summable_on UNIV \ (\x. -fds_nth f x * mangoldt x / nat_power x s) abs_summable_on Collect primepow" unfolding h_def using fds_nth_logderiv_completely_multiplicative[OF assms(2)] by (intro abs_summable_on_cong_neutral) (auto simp: fds_nth_fds mangoldt_def) finally have sum1: "(\x. -fds_nth f x * mangoldt x / nat_power x s) abs_summable_on Collect primepow" by (rule abs_summable_on_subset) auto also have "?this \ (\(p,k). -fds_nth f (p ^ Suc k) * mangoldt (p ^ Suc k) / nat_power (p ^ Suc k) s) abs_summable_on (?P \ UNIV)" using bij_betw_primepows unfolding case_prod_unfold by (intro abs_summable_on_reindex_bij_betw [symmetric]) also have "\ \ (\(p,k). -((fds_nth f p / nat_power p s) ^ Suc k * of_real (ln (real p)))) abs_summable_on (?P \ UNIV)" unfolding case_prod_unfold by (intro abs_summable_on_cong, subst mangoldt_primepow) (auto simp: f.mult f.power nat_power_mult_distrib nat_power_power_left power_divide dest: prime_gt_1_nat) finally have sum2: \ . have sum4: "summable (\n. (norm (fds_nth f p / nat_power p s)) ^ Suc n)" if p: "prime p" for p proof - have "summable (\n. \ln (real p)\ * (norm (fds_nth f p / nat_power p s)) ^ Suc n)" using p abs_summable_on_Sigma_project2[OF sum2, of p] unfolding abs_summable_on_nat_iff' by (simp add: norm_power norm_mult norm_divide mult_ac del: power_Suc) thus ?thesis by (rule summable_mult_D) (insert p, auto dest: prime_gt_1_nat) qed have sums: "(\n. (fds_nth f p / nat_power p s) ^ Suc n) sums (1 / (1 - fds_nth f p / nat_power p s) - 1)" if p: "prime p" for p :: nat proof - from sum4[OF p] have "norm (fds_nth f p / nat_power p s) < 1" unfolding summable_Suc_iff by (simp add: summable_geometric_iff) from geometric_sums[OF this] show ?thesis by (subst sums_Suc_iff) auto qed have eq: "(\\<^sub>ak. - ((fds_nth f p / nat_power p s) ^ Suc k * of_real (ln (real p)))) = -(of_real (ln (real p)) * (1 / (1 - fds_nth f p / nat_power p s) - 1))" if p: "prime p" for p proof - have "(\\<^sub>ak. - ((fds_nth f p / nat_power p s) ^ Suc k * of_real (ln (real p)))) = (\\<^sub>ak. (fds_nth f p / nat_power p s) ^ Suc k) * of_real (-ln (real p))" using sum4[of p] p by (subst infsetsum_cmult_left [symmetric]) (auto simp: abs_summable_on_nat_iff' norm_power simp del: power_Suc) also have "(\\<^sub>ak. (fds_nth f p / nat_power p s) ^ Suc k) = (1 / (1 - fds_nth f p / nat_power p s) - 1)" using sum4[OF p] sums[OF p] by (subst infsetsum_nat') (auto simp: sums_iff abs_summable_on_nat_iff' norm_power simp del: power_Suc) finally show ?thesis by (simp add: mult_ac) qed have sum3: "(\x. \\<^sub>ay. - ((fds_nth f x / nat_power x s) ^ Suc y * of_real (ln (real x)))) abs_summable_on {p. prime p}" using sum2 by (rule abs_summable_on_Sigma_project1') auto also have "?this \ (\p. -(of_real (ln (real p)) * (1 / (1 - fds_nth f p / nat_power p s) - 1))) abs_summable_on {p. prime p}" by (intro abs_summable_on_cong eq) auto also have "\ \ ?th1" by (subst abs_summable_on_uminus_iff) auto finally show ?th1 . have "eval_fds h s = (\\<^sub>an. fds_nth h n / nat_power n s)" using * unfolding eval_fds_def by (subst infsetsum_nat') auto also have "\ = (\\<^sub>an \ {n. primepow n}. -fds_nth f n * mangoldt n / nat_power n s)" unfolding h_def using fds_nth_logderiv_completely_multiplicative[OF assms(2)] by (intro infsetsum_cong_neutral) (auto simp: fds_nth_fds mangoldt_def) also have "\ = (\\<^sub>a(p,k)\(?P \ UNIV). -fds_nth f (p ^ Suc k) * mangoldt (p ^ Suc k) / nat_power (p ^ Suc k) s)" using bij_betw_primepows unfolding case_prod_unfold by (intro infsetsum_reindex_bij_betw [symmetric]) also have "\ = (\\<^sub>a(p,k)\(?P \ UNIV). -((fds_nth f p / nat_power p s) ^ Suc k) * of_real (ln (real p)))" by (intro infsetsum_cong) (auto simp: f.mult f.power mangoldt_def aprimedivisor_prime_power ln_realpow prime_gt_0_nat nat_power_power_left divide_simps simp del: power_Suc) also have "\ = (\\<^sub>ap | prime p. \\<^sub>ak. - ((fds_nth f p / nat_power p s) ^ Suc k) * of_real (ln (real p)))" using sum2 by (subst infsetsum_Times) (auto simp: case_prod_unfold) also have "\ = (\\<^sub>ap | prime p. -(of_real (ln (real p)) * (1 / (1 - fds_nth f p / nat_power p s) - 1)))" using eq by (intro infsetsum_cong) auto finally show ?th2 by (subst (asm) infsetsum_uminus) qed lemma eval_fds_logderiv_zeta: assumes "Re s > 1" shows "(\p. of_real (ln (real p)) / (p powr s - 1)) abs_summable_on {p. prime p}" (is ?th1) and "deriv zeta s / zeta s = -(\\<^sub>ap | prime p. of_real (ln (real p)) / (p powr s - 1))" (is ?th2) proof - have *: "completely_multiplicative_function (fds_nth fds_zeta :: _ \ complex)" by standard auto note abscissa = le_less_trans[OF abs_conv_abscissa_completely_multiplicative_log_deriv[OF *]] have "(\p. ln (real p) * (1 / (1 - fds_nth fds_zeta p / p powr s) - 1)) abs_summable_on {p. prime p}" using eval_fds_logderiv_completely_multiplicative[OF *, of s] assms by auto also have "?this \ (\p. ln (real p) / (p powr s - 1)) abs_summable_on {p. prime p}" using assms by (intro abs_summable_on_cong) (auto simp: fds_nth_zeta divide_simps dest: prime_gt_1_nat) finally show ?th1 . from assms have ev: "eventually (\z. z \ {z. Re z > 1}) (nhds s)" by (intro eventually_nhds_in_open open_halfspace_Re_gt) auto have "deriv zeta s = deriv (eval_fds fds_zeta) s" by (intro deriv_cong_ev[OF eventually_mono[OF ev]]) (auto simp: eval_fds_zeta) also have "deriv (eval_fds fds_zeta) s / zeta s = eval_fds (fds_deriv fds_zeta / fds_zeta) s" using assms zeta_Re_gt_1_nonzero[of s] by (subst eval_fds_log_deriv) (auto simp: eval_fds_zeta eval_fds_deriv intro!: abscissa) also have "eval_fds (fds_deriv fds_zeta / fds_zeta) s = -(\\<^sub>ap | prime p. ln (real p) * (1 / (1 - fds_nth fds_zeta p / p powr s) - 1))" (is "_ = -?S") using eval_fds_logderiv_completely_multiplicative[OF *, of s] assms by auto also have "?S = (\\<^sub>ap | prime p. ln (real p) / (p powr s - 1))" using assms by (intro infsetsum_cong) (auto simp: fds_nth_zeta divide_simps dest: prime_gt_1_nat) finally show ?th2 . qed lemma sums_logderiv_zeta: assumes "Re s > 1" shows "(\p. if prime p then of_real (ln (real p)) / (of_nat p powr s - 1) else 0) sums -(deriv zeta s / zeta s)" (is "?f sums _") proof - note * = eval_fds_logderiv_zeta[OF assms] from sums_infsetsum_nat[OF *(1)] and *(2) show ?thesis by simp qed lemma abs_conv_abscissa_diff_le: "abs_conv_abscissa (f - g :: 'a :: dirichlet_series fds) \ max (abs_conv_abscissa f) (abs_conv_abscissa g)" using abs_conv_abscissa_add_le[of f "-g"] by auto lemma abs_conv_abscissa_diff_leI: "abs_conv_abscissa (f :: 'a :: dirichlet_series fds) \ d \ abs_conv_abscissa g \ d \ abs_conv_abscissa (f - g) \ d" using abs_conv_abscissa_diff_le[of f g] by (auto simp: le_max_iff_disj) lemma range_add_nat: "range (\n. n + c) = {(c::nat)..}" proof safe fix x assume "x \ c" hence "x = x - c + c" by simp thus "x \ range (\n. n + c)" by blast qed auto lemma abs_summable_hurwitz_zeta: assumes "Re s > 1" "a + real b > 0" shows "(\n. 1 / (of_nat n + a) powr s) abs_summable_on {b..}" proof - from assms have "summable (\n. cmod (1 / (of_nat (n + b) + a) powr s))" using summable_hurwitz_zeta_real[of "Re s" "a + b"] by (auto simp: norm_divide powr_minus field_simps norm_powr_real_powr) hence "(\n. 1 / (of_nat (n + b) + a) powr s) abs_summable_on UNIV" by (auto simp: abs_summable_on_nat_iff' add_ac) also have "?this \ (\n. 1 / (of_nat n + a) powr s) abs_summable_on range (\n. n + b)" by (rule abs_summable_on_reindex_iff) auto also have "range (\n. n + b) = {b..}" by (rule range_add_nat) finally show ?thesis . qed lemma hurwitz_zeta_nat_conv_infsetsum: assumes "a > 0" and "Re s > 1" shows "hurwitz_zeta (real a) s = (\\<^sub>an. of_nat (n + a) powr -s)" "hurwitz_zeta (real a) s = (\\<^sub>an\{a..}. of_nat n powr -s)" proof - have "hurwitz_zeta (real a) s = (\n. of_nat (n + a) powr -s)" using assms by (subst hurwitz_zeta_conv_suminf) auto also have "\ = (\\<^sub>an. of_nat (n + a) powr -s)" using abs_summable_hurwitz_zeta[of s a 0] assms by (intro infsetsum_nat' [symmetric]) (auto simp: powr_minus field_simps) finally show "hurwitz_zeta (real a) s = (\\<^sub>an. of_nat (n + a) powr -s)" . also have "\ = (\\<^sub>an\range (\n. n + a). of_nat n powr -s)" by (rule infsetsum_reindex [symmetric]) auto also have "range (\n. n + a) = {a..}" by (rule range_add_nat) finally show "hurwitz_zeta (real a) s = (\\<^sub>an\{a..}. of_nat n powr -s)" . qed lemma continuous_on_pre_zeta [continuous_intros]: assumes "continuous_on A f" "a > 0" shows "continuous_on A (\x. pre_zeta a (f x))" proof - from assms have "continuous_on UNIV (pre_zeta a)" by (intro holomorphic_on_imp_continuous_on[OF holomorphic_pre_zeta]) auto from continuous_on_compose2[OF this assms(1)] show ?thesis by simp qed lemma continuous_pre_zeta [continuous_intros]: assumes "continuous (at x within A) f" "a > 0" shows "continuous (at x within A) (\x. pre_zeta a (f x))" proof - have "continuous (at z) (pre_zeta a)" for z by (rule continuous_on_interior[of UNIV]) (insert assms, auto intro!: continuous_intros) from continuous_within_compose3[OF this assms(1)] show ?thesis . qed lemma pre_zeta_bound: assumes "0 < Re s" and a: "a > 0" shows "norm (pre_zeta a s) \ (1 + norm s / Re s) / 2 * a powr -Re s" proof - let ?f = "\x. - (s * (x + a) powr (-1-s))" let ?g' = "\x. norm s * (x + a) powr (-1-Re s)" let ?g = "\x. -norm s / Re s * (x + a) powr (-Re s)" define R where "R = EM_remainder 1 ?f 0" have [simp]: "-Re s - 1 = -1 - Re s" by (simp add: algebra_simps) have "\frac x - 1 / 2\ \ 1 / 2" for x :: real unfolding frac_def by linarith hence "\pbernpoly (Suc 0) x\ \ 1 / 2" for x by (simp add: pbernpoly_def bernpoly_def) moreover have "((\b. cmod s * (b + a) powr - Re s / Re s) \ 0) at_top" using \Re s > 0\ \a > 0\ by real_asymp ultimately have *: "\x. x \ real 0 \ norm (EM_remainder 1 ?f (int x)) \ (1 / 2) / fact 1 * (-?g (real x))" using \a > 0\ \Re s > 0\ by (intro norm_EM_remainder_le_strong_nat'[where g' = ?g' and Y = "{}"]) (auto intro!: continuous_intros derivative_eq_intros simp: field_simps norm_mult norm_powr_real_powr add_eq_0_iff) have R: "norm R \ norm s / (2 * Re s) * a powr -Re s" unfolding R_def using spec[OF *, of 0] by simp from assms have "pre_zeta a s = a powr -s / 2 + R" by (simp add: pre_zeta_def pre_zeta_aux_def R_def) also have "norm \ \ a powr -Re s / 2 + norm s / (2 * Re s) * a powr -Re s" using a by (intro order.trans[OF norm_triangle_ineq] add_mono R) (auto simp: norm_powr_real_powr) also have "\ = (1 + norm s / Re s) / 2 * a powr -Re s" by (simp add: field_simps) finally show ?thesis . qed lemma pre_zeta_bound': assumes "0 < Re s" and a: "a > 0" shows "norm (pre_zeta a s) \ norm s / (Re s * a powr Re s)" proof - from assms have "norm (pre_zeta a s) \ (1 + norm s / Re s) / 2 * a powr -Re s" by (intro pre_zeta_bound) auto also have "\ = (Re s + norm s) / 2 / (Re s * a powr Re s)" using assms by (auto simp: field_simps powr_minus) also have "Re s + norm s \ norm s + norm s" by (intro add_right_mono complex_Re_le_cmod) also have "(norm s + norm s) / 2 = norm s" by simp finally show "norm (pre_zeta a s) \ norm s / (Re s * a powr Re s)" using assms by (simp add: divide_right_mono) qed lemma summable_comparison_test_bigo: fixes f :: "nat \ real" assumes "summable (\n. norm (g n))" "f \ O(g)" shows "summable f" proof - from \f \ O(g)\ obtain C where C: "eventually (\x. norm (f x) \ C * norm (g x)) at_top" by (auto elim: landau_o.bigE) thus ?thesis by (rule summable_comparison_test_ev) (insert assms, auto intro: summable_mult) qed lemma deriv_zeta_eq: assumes s: "s \ 1" shows "deriv zeta s = deriv (pre_zeta 1) s - 1 / (s - 1)\<^sup>2" proof - from s have ev: "eventually (\z. z \ 1) (nhds s)" by (intro t1_space_nhds) have [derivative_intros]: "(pre_zeta 1 has_field_derivative deriv (pre_zeta 1) s) (at s)" by (intro holomorphic_derivI[of _ UNIV] holomorphic_intros) auto have "((\s. pre_zeta 1 s + 1 / (s - 1)) has_field_derivative (deriv (pre_zeta 1) s - 1 / (s - 1)\<^sup>2)) (at s)" using s by (auto intro!: derivative_eq_intros simp: power2_eq_square) also have "?this \ (zeta has_field_derivative (deriv (pre_zeta 1) s - 1 / (s - 1)\<^sup>2)) (at s)" by (intro has_field_derivative_cong_ev eventually_mono[OF ev]) (auto simp: zeta_def hurwitz_zeta_def) finally show ?thesis by (rule DERIV_imp_deriv) qed lemma zeta_remove_zero: assumes "Re s \ 1" shows "(s - 1) * pre_zeta 1 s + 1 \ 0" proof (cases "s = 1") case False hence "(s - 1) * pre_zeta 1 s + 1 = (s - 1) * zeta s" by (simp add: zeta_def hurwitz_zeta_def divide_simps) also from False assms have "\ \ 0" using zeta_Re_ge_1_nonzero[of s] by auto finally show ?thesis . qed auto lemma eval_fds_deriv_zeta: assumes "Re s > 1" shows "eval_fds (fds_deriv fds_zeta) s = deriv zeta s" proof - have ev: "eventually (\z. z \ {z. Re z > 1}) (nhds s)" using assms by (intro eventually_nhds_in_open open_halfspace_Re_gt) auto from assms have "eval_fds (fds_deriv fds_zeta) s = deriv (eval_fds fds_zeta) s" by (subst eval_fds_deriv) auto also have "\ = deriv zeta s" by (intro deriv_cong_ev eventually_mono[OF ev]) (auto simp: eval_fds_zeta) finally show ?thesis . qed lemma length_sorted_list_of_set [simp]: "finite A \ length (sorted_list_of_set A) = card A" by (metis length_remdups_card_conv length_sort set_sorted_list_of_set sorted_list_of_set_sort_remdups) lemma le_nat_iff': "x \ nat y \ x = 0 \ y \ 0 \ int x \ y" by auto lemma sum_upto_plus1: assumes "x \ 0" shows "sum_upto f (x + 1) = sum_upto f x + f (Suc (nat \x\))" proof - have "sum_upto f (x + 1) = sum f {0<..Suc (nat \x\)}" using assms by (simp add: sum_upto_altdef nat_add_distrib) also have "{0<..Suc (nat \x\)} = insert (Suc (nat \x\)) {0<..nat \x\}" by auto also have "sum f \ = sum_upto f x + f (Suc (nat \x\))" by (subst sum.insert) (auto simp: sum_upto_altdef add_ac) finally show ?thesis . qed lemma sum_upto_minus1: assumes "x \ 1" shows "sum_upto f (x - 1) = (sum_upto f x - f (nat \x\) :: 'a :: ab_group_add)" using sum_upto_plus1[of "x - 1" f] assms by (simp add: algebra_simps nat_diff_distrib) lemma integral_smallo: fixes f g g' :: "real \ real" assumes "f \ o(g')" and "filterlim g at_top at_top" assumes "\a' x. a \ a' \ a' \ x \ f integrable_on {a'..x}" assumes deriv: "\x. x \ a \ (g has_field_derivative g' x) (at x)" assumes cont: "continuous_on {a..} g'" assumes nonneg: "\x. x \ a \ g' x \ 0" shows "(\x. integral {a..x} f) \ o(g)" proof (rule landau_o.smallI) fix c :: real assume c: "c > 0" note [continuous_intros] = continuous_on_subset[OF cont] define c' where "c' = c / 2" from c have c': "c' > 0" by (simp add: c'_def) from landau_o.smallD[OF assms(1) this] obtain b where b: "\x. x \ b \ norm (f x) \ c' * norm (g' x)" unfolding eventually_at_top_linorder by blast define b' where "b' = max a b" define D where "D = norm (integral {a..b'} f)" have "filterlim (\x. c' * g x) at_top at_top" using c' by (intro filterlim_tendsto_pos_mult_at_top[OF tendsto_const] assms) hence "eventually (\x. c' * g x \ D - c' * g b') at_top" by (auto simp: filterlim_at_top) thus "eventually (\x. norm (integral {a..x} f) \ c * norm (g x)) at_top" using eventually_ge_at_top[of b'] proof eventually_elim case (elim x) have b': "a \ b'" "b \ b'" by (auto simp: b'_def) from elim b' have integrable: "(\x. \g' x\) integrable_on {b'..x}" by (intro integrable_continuous_real continuous_intros) auto have "integral {a..x} f = integral {a..b'} f + integral {b'..x} f" using elim b' by (intro Henstock_Kurzweil_Integration.integral_combine [symmetric] assms) auto also have "norm \ \ D + norm (integral {b'..x} f)" unfolding D_def by (rule norm_triangle_ineq) also have "norm (integral {b'..x} f) \ integral {b'..x} (\x. c' * norm (g' x))" using b' elim assms c' integrable by (intro integral_norm_bound_integral b assms) auto also have "\ = c' * integral {b'..x} (\x. \g' x\)" by simp also have "integral {b'..x} (\x. \g' x\) = integral {b'..x} g'" using assms b' by (intro integral_cong) auto also have "(g' has_integral (g x - g b')) {b'..x}" using b' elim by (intro fundamental_theorem_of_calculus) (auto simp flip: has_field_derivative_iff_has_vector_derivative intro!: has_field_derivative_at_within[OF deriv]) hence "integral {b'..x} g' = g x - g b'" by (simp add: has_integral_iff) also have "D + c' * (g x - g b') \ c * g x" using elim by (simp add: field_simps c'_def) also have "\ \ c * norm (g x)" using c by (intro mult_left_mono) auto finally show ?case by simp qed qed lemma integral_bigo: fixes f g g' :: "real \ real" assumes "f \ O(g')" and "filterlim g at_top at_top" assumes "\a' x. a \ a' \ a' \ x \ f integrable_on {a'..x}" assumes deriv: "\x. x \ a \ (g has_field_derivative g' x) (at x within {a..})" assumes cont: "continuous_on {a..} g'" assumes nonneg: "\x. x \ a \ g' x \ 0" shows "(\x. integral {a..x} f) \ O(g)" proof - note [continuous_intros] = continuous_on_subset[OF cont] from landau_o.bigE[OF assms(1)] obtain c b where c: "c > 0" and b: "\x. x \ b \ norm (f x) \ c * norm (g' x)" unfolding eventually_at_top_linorder by metis define c' where "c' = c / 2" define b' where "b' = max a b" define D where "D = norm (integral {a..b'} f)" have "filterlim (\x. c * g x) at_top at_top" using c by (intro filterlim_tendsto_pos_mult_at_top[OF tendsto_const] assms) hence "eventually (\x. c * g x \ D - c * g b') at_top" by (auto simp: filterlim_at_top) hence "eventually (\x. norm (integral {a..x} f) \ 2 * c * norm (g x)) at_top" using eventually_ge_at_top[of b'] proof eventually_elim case (elim x) have b': "a \ b'" "b \ b'" by (auto simp: b'_def) from elim b' have integrable: "(\x. \g' x\) integrable_on {b'..x}" by (intro integrable_continuous_real continuous_intros) auto have "integral {a..x} f = integral {a..b'} f + integral {b'..x} f" using elim b' by (intro Henstock_Kurzweil_Integration.integral_combine [symmetric] assms) auto also have "norm \ \ D + norm (integral {b'..x} f)" unfolding D_def by (rule norm_triangle_ineq) also have "norm (integral {b'..x} f) \ integral {b'..x} (\x. c * norm (g' x))" using b' elim assms c integrable by (intro integral_norm_bound_integral b assms) auto also have "\ = c * integral {b'..x} (\x. \g' x\)" by simp also have "integral {b'..x} (\x. \g' x\) = integral {b'..x} g'" using assms b' by (intro integral_cong) auto also have "(g' has_integral (g x - g b')) {b'..x}" using b' elim by (intro fundamental_theorem_of_calculus) (auto simp flip: has_field_derivative_iff_has_vector_derivative intro!: DERIV_subset[OF deriv]) hence "integral {b'..x} g' = g x - g b'" by (simp add: has_integral_iff) also have "D + c * (g x - g b') \ 2 * c * g x" using elim by (simp add: field_simps c'_def) also have "\ \ 2 * c * norm (g x)" using c by (intro mult_left_mono) auto finally show ?case by simp qed thus ?thesis by (rule bigoI) qed lemma primepows_le_subset: assumes x: "x > 0" and l: "l > 0" shows "{(p, i). prime p \ l \ i \ real (p ^ i) \ x} \ {..nat \root l x\} \ {..nat \log 2 x\}" proof safe fix p i :: nat assume pi: "prime p" "i \ l" "real (p ^ i) \ x" have "real p ^ l \ real p ^ i" using pi x l by (intro power_increasing) (auto dest: prime_gt_0_nat) also have "\ \ x" using pi by simp finally have "root l (real p ^ l) \ root l x" using x pi l by (subst real_root_le_iff) auto also have "root l (real p ^ l) = real p" using pi l by (subst real_root_pos2) auto finally show "p \ nat \root l x\" using pi l x by (simp add: le_nat_iff' le_floor_iff) from pi have "2 ^ i \ real p ^ i" using l by (intro power_mono) (auto dest: prime_gt_1_nat) also have "\ \ x" using pi by simp finally show "i \ nat \log 2 x\" using pi x by (auto simp: le_nat_iff' le_floor_iff le_log_iff powr_realpow) qed lemma mangoldt_non_primepow: "\primepow n \ mangoldt n = 0" by (auto simp: mangoldt_def) lemma le_imp_bigo_real: assumes "c \ 0" "eventually (\x. f x \ c * (g x :: real)) F" "eventually (\x. 0 \ f x) F" shows "f \ O[F](g)" proof - have "eventually (\x. norm (f x) \ c * norm (g x)) F" using assms(2,3) proof eventually_elim case (elim x) have "norm (f x) \ c * g x" using elim by simp also have "\ \ c * norm (g x)" by (intro mult_left_mono assms) auto finally show ?case . qed thus ?thesis by (intro bigoI[of _ c]) auto qed (* TODO: unneeded. But why does real_asymp not work? *) lemma ln_minus_ln_floor_bigo: "(\x. ln x - ln (real (nat \x\))) \ O(\_. 1)" proof (intro le_imp_bigo_real[of 1] eventually_mono[OF eventually_ge_at_top[of 1]]) fix x :: real assume x: "x \ 1" from x have *: "x - real (nat \x\) \ 1" by linarith from x have "ln x - ln (real (nat \x\)) \ (x - real (nat \x\)) / real (nat \x\)" by (intro ln_diff_le) auto also have "\ \ 1 / 1" using x * by (intro frac_le) auto finally show "ln x - ln (real (nat \x\)) \ 1 * 1" by simp qed auto lemma cos_geD: assumes "cos x \ cos a" "0 \ a" "a \ pi" "-pi \ x" "x \ pi" shows "x \ {-a..a}" proof (cases "x \ 0") case True with assms show ?thesis by (subst (asm) cos_mono_le_eq) auto next case False with assms show ?thesis using cos_mono_le_eq[of a "-x"] by auto qed (* TODO: Could be generalised *) lemma path_image_part_circlepath_same_Re: assumes "0 \ b" "b \ pi" "a = -b" "r \ 0" shows "path_image (part_circlepath c r a b) = sphere c r \ {s. Re s \ Re c + r * cos a}" proof safe fix z assume "z \ path_image (part_circlepath c r a b)" with assms obtain t where t: "t \ {a..b}" "z = c + of_real r * cis t" by (auto simp: path_image_part_circlepath exp_eq_polar) from t and assms show "z \ sphere c r" by (auto simp: dist_norm norm_mult) from t and assms show "Re z \ Re c + r * cos a" using cos_monotone_0_pi_le[of t b] cos_monotone_minus_pi_0'[of a t] by (cases "t \ 0") (auto intro!: mult_left_mono) next fix z assume z: "z \ sphere c r" "Re z \ Re c + r * cos a" show "z \ path_image (part_circlepath c r a b)" proof (cases "r = 0") case False with assms have r: "r > 0" by simp with z have z_eq: "z = c + r * cis (Arg (z - c))" using Arg_eq[of "z - c"] by (auto simp: dist_norm exp_eq_polar norm_minus_commute) moreover from z(2) r assms have "cos b \ cos (Arg (z - c))" by (subst (asm) z_eq) auto with assms have "Arg (z - c) \ {-b..b}" using Arg_le_pi[of "z - c"] mpi_less_Arg[of "z - c"] by (intro cos_geD) auto ultimately show "z \ path_image (part_circlepath c r a b)" using assms by (subst path_image_part_circlepath) (auto simp: exp_eq_polar) qed (insert assms z, auto simp: path_image_part_circlepath) qed lemma part_circlepath_rotate_left: "part_circlepath c r (x + a) (x + b) = (\z. c + cis x * (z - c)) \ part_circlepath c r a b" by (simp add: part_circlepath_def exp_eq_polar fun_eq_iff linepath_translate_left linepath_translate_right cis_mult add_ac) lemma part_circlepath_rotate_right: "part_circlepath c r (a + x) (b + x) = (\z. c + cis x * (z - c)) \ part_circlepath c r a b" by (simp add: part_circlepath_def exp_eq_polar fun_eq_iff linepath_translate_left linepath_translate_right cis_mult add_ac) lemma path_image_semicircle_Re_ge: assumes "r \ 0" shows "path_image (part_circlepath c r (-pi/2) (pi/2)) = sphere c r \ {s. Re s \ Re c}" by (subst path_image_part_circlepath_same_Re) (simp_all add: assms) lemma sphere_rotate: "(\z. c + cis x * (z - c)) ` sphere c r = sphere c r" proof safe fix z assume z: "z \ sphere c r" hence "z = c + cis x * (c + cis (-x) * (z - c) - c)" "c + cis (-x) * (z - c) \ sphere c r" by (auto simp: dist_norm norm_mult norm_minus_commute cis_conv_exp exp_minus field_simps norm_divide) with z show "z \ (\z. c + cis x * (z - c)) ` sphere c r" by blast qed (auto simp: dist_norm norm_minus_commute norm_mult) lemma path_image_semicircle_Re_le: assumes "r \ 0" shows "path_image (part_circlepath c r (pi/2) (3/2*pi)) = sphere c r \ {s. Re s \ Re c}" proof - let ?f = "(\z. c + cis pi * (z - c))" have *: "part_circlepath c r (pi/2) (3/2*pi) = part_circlepath c r (pi + (-pi/2)) (pi + pi/2)" by simp have "path_image (part_circlepath c r (pi/2) (3/2*pi)) = ?f ` sphere c r \ ?f ` {s. Re c \ Re s}" unfolding * part_circlepath_rotate_left path_image_compose path_image_semicircle_Re_ge[OF assms] by auto also have "?f ` sphere c r = sphere c r" by (rule sphere_rotate) also have "?f ` {s. Re c \ Re s} = {s. Re c \ Re s}" by (auto simp: image_iff intro!: exI[of _ "2 * c - x" for x]) finally show ?thesis . qed lemma path_image_semicircle_Im_ge: assumes "r \ 0" shows "path_image (part_circlepath c r 0 pi) = sphere c r \ {s. Im s \ Im c}" proof - let ?f = "(\z. c + cis (pi/2) * (z - c))" have *: "part_circlepath c r 0 pi = part_circlepath c r (pi / 2 + (-pi/2)) (pi / 2 + pi/2)" by simp have "path_image (part_circlepath c r 0 pi) = ?f ` sphere c r \ ?f ` {s. Re c \ Re s}" unfolding * part_circlepath_rotate_left path_image_compose path_image_semicircle_Re_ge[OF assms] by auto also have "?f ` sphere c r = sphere c r" by (rule sphere_rotate) also have "?f ` {s. Re c \ Re s} = {s. Im c \ Im s}" by (auto simp: image_iff intro!: exI[of _ "c - \ * (x - c)" for x]) finally show ?thesis . qed lemma path_image_semicircle_Im_le: assumes "r \ 0" shows "path_image (part_circlepath c r pi (2 * pi)) = sphere c r \ {s. Im s \ Im c}" proof - let ?f = "(\z. c + cis (3*pi/2) * (z - c))" have *: "part_circlepath c r pi (2*pi) = part_circlepath c r (3*pi/2 + (-pi/2)) (3*pi/2 + pi/2)" by simp have "path_image (part_circlepath c r pi (2 * pi)) = ?f ` sphere c r \ ?f ` {s. Re c \ Re s}" unfolding * part_circlepath_rotate_left path_image_compose path_image_semicircle_Re_ge[OF assms] by auto also have "?f ` sphere c r = sphere c r" by (rule sphere_rotate) also have "cis (3 * pi / 2) = -\" using cis_mult[of pi "pi / 2"] by simp hence "?f ` {s. Re c \ Re s} = {s. Im c \ Im s}" by (auto simp: image_iff intro!: exI[of _ "c + \ * (x - c)" for x]) finally show ?thesis . qed lemma powr_numeral [simp]: "x \ 0 \ (x::real) powr numeral y = x ^ numeral y" using powr_numeral[of x y] by (cases "x = 0") auto lemma eval_fds_logderiv_zeta_real: assumes "x > (1 :: real)" shows "(\p. ln (real p) / (p powr x - 1)) abs_summable_on {p. prime p}" (is ?th1) and "deriv zeta (of_real x) / zeta (of_real x) = -of_real (\\<^sub>ap | prime p. ln (real p) / (p powr x - 1))" (is ?th2) proof - have "(\p. Re (of_real (ln (real p)) / (of_nat p powr of_real x - 1))) abs_summable_on {p. prime p}" using assms by (intro abs_summable_Re eval_fds_logderiv_zeta) auto also have "?this \ ?th1" by (intro abs_summable_on_cong) (auto simp: powr_Reals_eq) finally show ?th1 . show ?th2 using assms by (subst eval_fds_logderiv_zeta) (auto simp: infsetsum_of_real [symmetric] powr_Reals_eq) qed lemma fixes a b c d :: real assumes ab: "d * a + b \ 1" and c: "c < -1" and d: "d > 0" defines "C \ - ((ln (d * a + b) - 1 / (c + 1)) * (d * a + b) powr (c + 1) / (d * (c + 1)))" shows set_integrable_ln_powr_at_top: "(\x. (ln (d * x + b) * ((d * x + b) powr c))) absolutely_integrable_on {a<..}" (is ?th1) and set_lebesgue_integral_ln_powr_at_top: "(\x\{a<..}. (ln (d * x + b) * ((d * x + b) powr c)) \lborel) = C" (is ?th2) and ln_powr_has_integral_at_top: "((\x. ln (d * x + b) * (d * x + b) powr c) has_integral C) {a<..}" (is ?th3) proof - define f where "f = (\x. ln (d * x + b) * (d * x + b) powr c)" define F where "F = (\x. (ln (d * x + b) - 1 / (c + 1)) * (d * x + b) powr (c + 1) / (d * (c + 1)))" have *: "(F has_field_derivative f x) (at x)" "isCont f x" "f x \ 0" if "x > a" for x proof - have "1 \ d * a + b" by fact also have "\ < d * x + b" using that assms by (intro add_strict_right_mono mult_strict_left_mono) finally have gt_1: "d * x + b > 1" . show "(F has_field_derivative f x) (at x)" "isCont f x" using ab c d gt_1 by (auto simp: F_def f_def divide_simps intro!: derivative_eq_intros continuous_intros) (auto simp: algebra_simps powr_add)? show "f x \ 0" using gt_1 by (auto simp: f_def) qed have limits: "((F \ real_of_ereal) \ F a) (at_right (ereal a))" "((F \ real_of_ereal) \ 0) (at_left \)" using c ab d unfolding ereal_tendsto_simps1 F_def by (real_asymp; simp add: field_simps)+ have 1: "set_integrable lborel (einterval a \) f" using ab c limits by (intro interval_integral_FTC_nonneg) (auto intro!: * AE_I2) thus 2: "f absolutely_integrable_on {a<..}" by (auto simp: set_integrable_def integrable_completion) have "(LBINT x=ereal a..\. f x) = 0 - F a" using ab c limits by (intro interval_integral_FTC_nonneg) (auto intro!: *) thus 3: ?th2 by (simp add: interval_integral_to_infinity_eq F_def f_def C_def) show ?th3 using set_borel_integral_eq_integral[OF 1] 3 by (simp add: has_integral_iff f_def C_def) qed lemma ln_fact_conv_sum_upto: "ln (fact n) = sum_upto ln n" by (induction n) (auto simp: sum_upto_plus1 add.commute[of 1] ln_mult) lemma sum_upto_ln_conv_ln_fact: "sum_upto ln x = ln (fact (nat \x\))" by (simp add: ln_fact_conv_sum_upto sum_upto_altdef) lemma real_of_nat_div: "real (a div b) = real_of_int \real a / real b\" by (subst floor_divide_of_nat_eq) auto lemma integral_subset_negligible: fixes f :: "'a :: euclidean_space \ 'b :: banach" assumes "S \ T" "negligible (T - S)" shows "integral S f = integral T f" proof - have "integral T f = integral T (\x. if x \ S then f x else 0)" by (rule integral_spike[of "T - S"]) (use assms in auto) also have "\ = integral (S \ T) f" by (subst integral_restrict_Int) auto also have "S \ T = S" using assms by auto finally show ?thesis .. qed lemma integrable_on_cong [cong]: assumes "\x. x \ A \ f x = g x" "A = B" shows "f integrable_on A \ g integrable_on B" using has_integral_cong[of A f g, OF assms(1)] assms(2) by (auto simp: integrable_on_def) lemma measurable_sum_upto [measurable]: fixes f :: "'a \ nat \ real" assumes [measurable]: "\y. (\t. f t y) \ M \\<^sub>M borel" assumes [measurable]: "x \ M \\<^sub>M borel" shows "(\t. sum_upto (f t) (x t)) \ M \\<^sub>M borel" proof - have meas: "(\t. set_lebesgue_integral lborel {y. y \ 0 \ y - real (nat \x t\) \ 0} (\y. f t (nat \y\))) \ M \\<^sub>M borel" (is "?f \ _") unfolding set_lebesgue_integral_def by measurable also have "?f = (\t. sum_upto (f t) (x t))" proof fix t :: 'a show "?f t = sum_upto (f t) (x t)" proof (cases "x t < 1") case True hence "{y. y \ 0 \ y - real (nat \x t\) \ 0} = {0}" by auto thus ?thesis using True by (simp add: set_integral_at_point sum_upto_altdef) next case False define n where "n = nat \x t\" from False have "n > 0" by (auto simp: n_def) have *: "((\x. f t (nat \x\)) has_integral sum (f t) {0<..n}) {real 0..real n}" using \n > 0\ by (intro nat_sum_has_integral_ceiling) auto have **: "(\x. f t (nat \x\)) absolutely_integrable_on {real 0..real n}" proof (rule absolutely_integrable_absolutely_integrable_ubound) show "(\_. MAX n\{0..n}. \f t n\) absolutely_integrable_on {real 0..real n}" using \n > 0\ by (subst absolutely_integrable_on_iff_nonneg) (auto simp: Max_ge_iff intro!: exI[of _ "f t 0"]) show "(\x. f t (nat \x\)) integrable_on {real 0..real n}" using * by (simp add: has_integral_iff) next fix y :: real assume y: "y \ {real 0..real n}" have "f t (nat \y\) \ \f t (nat \y\)\" by simp also have "\ \ (MAX n\{0..n}. \f t n\)" using y by (intro Max.coboundedI) auto finally show "f t (nat \y\) \ (MAX n\{0..n}. \f t n\)" . qed have "sum (f t) {0<..n} = (\x\{real 0..real n}. f t (nat \x\) \lebesgue)" using has_integral_set_lebesgue[OF **] * by (simp add: has_integral_iff) also have "\ = (\x\{real 0..real n}. f t (nat \x\) \lborel)" unfolding set_lebesgue_integral_def by (subst integral_completion) auto also have "{real 0..real n} = {y. 0 \ y \ y - real (nat \x t\) \ 0}" by (auto simp: n_def) also have "sum (f t) {0<..n} = sum_upto (f t) (x t)" by (simp add: sum_upto_altdef n_def) finally show ?thesis .. qed qed finally show ?thesis . qed end