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/Counterclockwise_2D_Arbitrary.thy b/thys/Affine_Arithmetic/Counterclockwise_2D_Arbitrary.thy --- a/thys/Affine_Arithmetic/Counterclockwise_2D_Arbitrary.thy +++ b/thys/Affine_Arithmetic/Counterclockwise_2D_Arbitrary.thy @@ -1,939 +1,939 @@ section \CCW for Arbitrary Points in the Plane\ theory Counterclockwise_2D_Arbitrary imports Counterclockwise_2D_Strict begin subsection \Interpretation of Knuth's axioms in the plane\ definition lex::"point \ point \ bool" where "lex p q \ (fst p < fst q \ fst p = fst q \ snd p < snd q \ p = q)" definition psi::"point \ point \ point \ bool" where "psi p q r \ (lex p q \ lex q r)" definition ccw::"point \ point \ point \ bool" where "ccw p q r \ ccw' p q r \ (det3 p q r = 0 \ (psi p q r \ psi q r p \ psi r p q))" interpretation ccw: linorder_list0 "ccw x" for x . lemma ccw'_imp_ccw: "ccw' a b c \ ccw a b c" by (simp add: ccw_def) lemma ccw_ncoll_imp_ccw: "ccw a b c \ \coll a b c \ ccw' a b c" by (simp add: ccw_def) lemma ccw_translate: "ccw p (p + q) (p + r) = ccw 0 q r" by (auto simp: ccw_def psi_def lex_def) lemma ccw_translate_origin: "NO_MATCH 0 p \ ccw p q r = ccw 0 (q - p) (r - p)" using ccw_translate[of p "q - p" "r - p"] by simp lemma psi_scale: "psi (r *\<^sub>R a) (r *\<^sub>R b) 0 = (if r > 0 then psi a b 0 else if r < 0 then psi 0 b a else True)" "psi (r *\<^sub>R a) 0 (r *\<^sub>R b) = (if r > 0 then psi a 0 b else if r < 0 then psi b 0 a else True)" "psi 0 (r *\<^sub>R a) (r *\<^sub>R b) = (if r > 0 then psi 0 a b else if r < 0 then psi b a 0 else True)" - by (auto simp: psi_def lex_def det3_def' not_less sign_simps) + by (auto simp: psi_def lex_def det3_def' not_less algebra_split_simps) lemma ccw_scale23: "ccw 0 a b \ r > 0 \ ccw 0 (r *\<^sub>R a) (r *\<^sub>R b)" by (auto simp: ccw_def psi_scale) lemma psi_notI: "distinct3 p q r \ psi p q r \ \ psi q p r" by (auto simp: algebra_simps psi_def lex_def) lemma not_lex_eq: "\ lex a b \ lex b a \ a \ b" by (auto simp: algebra_simps lex_def prod_eq_iff) lemma lex_trans: "lex a b \ lex b c \ lex a c" by (auto simp: lex_def) lemma lex_sym_eqI: "lex a b \ lex b a \ a = b" and lex_sym_eq_iff: "lex a b \ lex b a \ a = b" by (auto simp: lex_def) lemma lex_refl[simp]: "lex p p" by (metis not_lex_eq) lemma psi_disjuncts: "distinct3 p q r \ psi p q r \ psi p r q \ psi q r p \ psi q p r \ psi r p q \ psi r q p" by (auto simp: psi_def not_lex_eq) lemma nlex_ccw_left: "lex x 0 \ ccw 0 (0, 1) x" by (auto simp: ccw_def lex_def psi_def ccw'_def det3_def') interpretation ccw_system123 ccw apply unfold_locales subgoal by (force simp: ccw_def ccw'_def det3_def' algebra_simps) subgoal by (force simp: ccw_def ccw'_def det3_def' psi_def algebra_simps lex_sym_eq_iff) subgoal by (drule psi_disjuncts) (force simp: ccw_def ccw'_def det3_def' algebra_simps) done lemma lex_scaleR_nonneg: "lex a b \ r \ 0 \ lex a (a + r *\<^sub>R (b - a))" by (auto simp: lex_def) lemma lex_scale1_zero: "lex (v *\<^sub>R u) 0 = (if v > 0 then lex u 0 else if v < 0 then lex 0 u else True)" and lex_scale2_zero: "lex 0 (v *\<^sub>R u) = (if v > 0 then lex 0 u else if v < 0 then lex u 0 else True)" - by (auto simp: lex_def prod_eq_iff less_eq_prod_def sign_simps) + by (auto simp: lex_def prod_eq_iff less_eq_prod_def algebra_split_simps) lemma nlex_add: assumes "lex a 0" "lex b 0" shows "lex (a + b) 0" using assms by (auto simp: lex_def) lemma nlex_sum: assumes "finite X" assumes "\x. x \ X \ lex (f x) 0" shows "lex (sum f X) 0" using assms by induction (auto intro!: nlex_add) lemma abs_add_nlex: assumes "coll 0 a b" assumes "lex a 0" assumes "lex b 0" shows "abs (a + b) = abs a + abs b" proof (rule antisym[OF abs_triangle_ineq]) have "fst (\a\ + \b\) \ fst \a + b\" using assms by (auto simp add: det3_def' abs_prod_def lex_def) moreover { assume H: "fst a < 0" "fst b < 0" hence "snd b \ 0 \ snd a \ 0" using assms by (auto simp: lex_def det3_def' mult.commute) (metis mult_le_cancel_left_neg mult_zero_right)+ hence "\snd a\ + \snd b\ \ \snd a + snd b\" using H by auto } hence "snd (\a\ + \b\) \ snd \a + b\" using assms by (auto simp add: det3_def' abs_prod_def lex_def) ultimately show "\a\ + \b\ \ \a + b\" unfolding less_eq_prod_def .. qed lemma lex_sum_list: "(\x. x \ set xs \ lex x 0) \ lex (sum_list xs) 0" by (induct xs) (auto simp: nlex_add) lemma abs_sum_list_coll: assumes coll: "list_all (coll 0 x) xs" assumes "x \ 0" assumes up: "list_all (\x. lex x 0) xs" shows "abs (sum_list xs) = sum_list (map abs xs)" using assms proof (induct xs) case (Cons y ys) hence "coll 0 x y" "coll 0 x (sum_list ys)" by (auto simp: list_all_iff intro!: coll_sum_list) hence "coll 0 y (sum_list ys)" using \x \ 0\ by (rule coll_trans) hence "\y + sum_list ys\ = abs y + abs (sum_list ys)" using Cons by (subst abs_add_nlex) (auto simp: list_all_iff lex_sum_list) thus ?case using Cons by simp qed simp lemma lex_diff1: "lex (a - b) c = lex a (c + b)" and lex_diff2: "lex c (a - b) = lex (c + b) a" by (auto simp: lex_def) lemma sum_list_eq_0_iff_nonpos: fixes xs::"'a::ordered_ab_group_add list" shows "list_all (\x. x \ 0) xs \ sum_list xs = 0 \ (\n\set xs. n = 0)" by (auto simp: list_all_iff sum_list_sum_nth sum_nonpos_eq_0_iff) (auto simp add: in_set_conv_nth) lemma sum_list_nlex_eq_zeroI: assumes nlex: "list_all (\x. lex x 0) xs" assumes "sum_list xs = 0" assumes "x \ set xs" shows "x = 0" proof - from assms(2) have z1: "sum_list (map fst xs) = 0" and z2: "sum_list (map snd xs) = 0" by (auto simp: prod_eq_iff fst_sum_list snd_sum_list) from nlex have "list_all (\x. x \ 0) (map fst xs)" by (auto simp: lex_def list_all_iff) from sum_list_eq_0_iff_nonpos[OF this] z1 nlex have z1': "list_all (\x. x = 0) (map fst xs)" and "list_all (\x. x \ 0) (map snd xs)" by (auto simp: list_all_iff lex_def) from sum_list_eq_0_iff_nonpos[OF this(2)] z2 have "list_all (\x. x = 0) (map snd xs)" by (simp add: list_all_iff) with z1' show "x = 0" by (auto simp: list_all_iff zero_prod_def assms prod_eq_iff) qed lemma sum_list_eq0I: "(\x\set xs. x = 0) \ sum_list xs = 0" by (induct xs) auto lemma sum_list_nlex_eq_zero_iff: assumes nlex: "list_all (\x. lex x 0) xs" shows "sum_list xs = 0 \ list_all ((=) 0) xs" using assms by (auto intro: sum_list_nlex_eq_zeroI sum_list_eq0I simp: list_all_iff) lemma assumes "lex p q" "lex q r" "0 \ a" "0 \ b" "0 \ c" "a + b + c = 1" assumes comb_def: "comb = a *\<^sub>R p + b *\<^sub>R q + c *\<^sub>R r" shows lex_convex3: "lex p comb" "lex comb r" proof - from convex3_alt[OF assms(3-6), of p q r] obtain u v where uv: "a *\<^sub>R p + b *\<^sub>R q + c *\<^sub>R r = p + u *\<^sub>R (q - p) + v *\<^sub>R (r - p)" "0 \ u" "0 \ v" "u + v \ 1" . have "lex p r" using assms by (metis lex_trans) hence "lex (v *\<^sub>R (p - r)) 0" using uv by (simp add: lex_scale1_zero lex_diff1) also have "lex 0 (u *\<^sub>R (q - p))" using \lex p q\ uv by (simp add: lex_scale2_zero lex_diff2) finally (lex_trans) show "lex p comb" unfolding comb_def uv by (simp add: lex_def prod_eq_iff algebra_simps) from comb_def have comb_def': "comb = c *\<^sub>R r + b *\<^sub>R q + a *\<^sub>R p" by simp from assms have "c + b + a = 1" by simp from convex3_alt[OF assms(5,4,3) this, of r q p] obtain u v where uv: "c *\<^sub>R r + b *\<^sub>R q + a *\<^sub>R p = r + u *\<^sub>R (q - r) + v *\<^sub>R (p - r)" "0 \ u" "0 \ v" "u + v \ 1" by auto have "lex (u *\<^sub>R (q - r)) 0" using uv \lex q r\ by (simp add: lex_scale1_zero lex_diff1) also have "lex 0 (v *\<^sub>R (r - p))" using uv \lex p r\ by (simp add: lex_scale2_zero lex_diff2) finally (lex_trans) show "lex comb r" unfolding comb_def' uv by (simp add: lex_def prod_eq_iff algebra_simps) qed lemma lex_convex_self2: assumes "lex p q" "0 \ a" "a \ 1" defines "r \ a *\<^sub>R p + (1 - a) *\<^sub>R q" shows "lex p r" (is ?th1) and "lex r q" (is ?th2) using lex_convex3[OF \lex p q\, of q a "1 - a" 0 r] assms by (simp_all add: r_def) lemma lex_uminus0[simp]: "lex (-a) 0 = lex 0 a" by (auto simp: lex_def) lemma lex_fst_zero_imp: "fst x = 0 \ lex x 0 \ lex y 0 \ \coll 0 x y \ ccw' 0 y x" - by (auto simp: ccw'_def det3_def' lex_def sign_simps) + by (auto simp: ccw'_def det3_def' lex_def algebra_split_simps) lemma lex_ccw_left: "lex x y \ r > 0 \ ccw y (y + (0, r)) x" by (auto simp: ccw_def ccw'_def det3_def' algebra_simps lex_def psi_def) lemma lex_translate_origin: "NO_MATCH 0 a \ lex a b = lex 0 (b - a)" by (auto simp: lex_def) subsection \Order prover setup\ definition "lexs p q \ (lex p q \ p \ q)" lemma lexs_irrefl: "\ lexs p p" and lexs_imp_lex: "lexs x y \ lex x y" and not_lexs: "(\ lexs x y) = (lex y x)" and not_lex: "(\ lex x y) = (lexs y x)" and eq_lex_refl: "x = y \ lex x y" by (auto simp: lexs_def lex_def prod_eq_iff) lemma lexs_trans: "lexs x y \ lexs y z \ lexs x z" and lexs_lex_trans: "lexs x y \ lex y z \ lexs x z" and lex_lexs_trans: "lex x y \ lexs y z \ lexs x z" and lex_neq_trans: "lex a b \ a \ b \ lexs a b" and neq_lex_trans: "a \ b \ lex a b \ lexs a b" and lexs_imp_neq: "lexs a b \ a \ b" by (auto simp: lexs_def lex_def prod_eq_iff) declare lexs_irrefl[THEN notE, order add less_reflE: linorder "(=) :: point => point => bool" lex lexs] declare lex_refl[order add le_refl: linorder "(=) :: point => point => bool" lex lexs] declare lexs_imp_lex[order add less_imp_le: linorder "(=) :: point => point => bool" lex lexs] declare not_lexs[THEN iffD2, order add not_lessI: linorder "(=) :: point => point => bool" lex lexs] declare not_lex[THEN iffD2, order add not_leI: linorder "(=) :: point => point => bool" lex lexs] declare not_lexs[THEN iffD1, order add not_lessD: linorder "(=) :: point => point => bool" lex lexs] declare not_lex[THEN iffD1, order add not_leD: linorder "(=) :: point => point => bool" lex lexs] declare lex_sym_eqI[order add eqI: linorder "(=) :: point => point => bool" lex lexs] declare eq_lex_refl[order add eqD1: linorder "(=) :: point => point => bool" lex lexs] declare sym[THEN eq_lex_refl, order add eqD2: linorder "(=) :: point => point => bool" lex lexs] declare lexs_trans[order add less_trans: linorder "(=) :: point => point => bool" lex lexs] declare lexs_lex_trans[order add less_le_trans: linorder "(=) :: point => point => bool" lex lexs] declare lex_lexs_trans[order add le_less_trans: linorder "(=) :: point => point => bool" lex lexs] declare lex_trans[order add le_trans: linorder "(=) :: point => point => bool" lex lexs] declare lex_neq_trans[order add le_neq_trans: linorder "(=) :: point => point => bool" lex lexs] declare neq_lex_trans[order add neq_le_trans: linorder "(=) :: point => point => bool" lex lexs] declare lexs_imp_neq[order add less_imp_neq: linorder "(=) :: point => point => bool" lex lexs] declare eq_neq_eq_imp_neq[order add eq_neq_eq_imp_neq: linorder "(=) :: point => point => bool" lex lexs] declare not_sym[order add not_sym: linorder "(=) :: point => point => bool" lex lexs] subsection \Contradictions\ lemma assumes d: "distinct4 s p q r" shows contra1: "\(lex p q \ lex q r \ lex r s \ indelta s p q r)" (is ?th1) and contra2: "\(lex s p \ lex p q \ lex q r \ indelta s p q r)" (is ?th2) and contra3: "\(lex p r \ lex p s \ lex q r \ lex q s \ insquare p r q s)" (is ?th3) proof - { assume "det3 s p q = 0" "det3 s q r = 0" "det3 s r p = 0" "det3 p q r = 0" hence ?th1 ?th2 ?th3 using d by (auto simp add: det3_def' ccw'_def ccw_def psi_def algebra_simps) } moreover { assume *: "\(det3 s p q = 0 \ det3 s q r = 0 \ det3 s r p = 0 \ det3 p q r = 0)" { assume d0: "det3 p q r = 0" with d have "?th1 \ ?th2" by (force simp add: det3_def' ccw'_def ccw_def psi_def algebra_simps) } moreover { assume dp: "det3 p q r \ 0" have "?th1 \ ?th2" unfolding de_Morgan_disj[symmetric] proof (rule notI, goal_cases) case prems: 1 hence **: "indelta s p q r" by auto hence nonnegs: "det3 p q r \ 0" "0 \ det3 s q r" "0 \ det3 p s r" "0 \ det3 p q s" by (auto simp: ccw_def ccw'_def det3_def' algebra_simps) hence det_pos: "det3 p q r > 0" using dp by simp have det_eq: "det3 s q r + det3 p s r + det3 p q s = det3 p q r" by (auto simp: ccw_def det3_def' algebra_simps) hence det_div_eq: "det3 s q r / det3 p q r + det3 p s r / det3 p q r + det3 p q s / det3 p q r = 1" using det_pos by (auto simp: field_simps) from lex_convex3[OF _ _ _ _ _ det_div_eq convex_comb_dets[OF det_pos, of s]] have "lex p s" "lex s r" using prems by (auto simp: nonnegs) with prems d show False by (simp add: lex_sym_eq_iff) qed } moreover have ?th3 proof (safe, goal_cases) case prems: 1 have nonnegs: "det3 p r q \ 0" "det3 r q s \ 0" "det3 s p r \ 0" "det3 q s p \ 0" using prems by (auto simp add: ccw_def ccw'_def less_eq_real_def) have dets_eq: "det3 p r q + det3 q s p = det3 r q s + det3 s p r" by (auto simp: det3_def') hence **: "det3 p r q = 0 \ det3 q s p = 0 \ det3 r q s = 0 \ det3 s p r = 0" using prems by (auto simp: ccw_def ccw'_def) moreover { fix p r q s assume det_pos: "det3 p r q > 0" assume dets_eq: "det3 p r q + det3 q s p = det3 r q s + det3 s p r" assume nonnegs:"det3 r q s \ 0" "det3 s p r \ 0" "det3 q s p \ 0" assume g14: "lex p r" "lex p s" "lex q r" "lex q s" assume d: "distinct4 s p q r" let ?sum = "(det3 p r q + det3 q s p) / det3 p r q" have eqs: "det3 s p r = det3 p r s" "det3 r q s = det3 s r q" "det3 q s p = - det3 p s q" by (auto simp: det3_def' algebra_simps) from convex_comb_dets[OF det_pos, of s] have "((det3 p r q / det3 p r q) *\<^sub>R s + (det3 q s p / det3 p r q) *\<^sub>R r) /\<^sub>R ?sum = ((det3 r q s / det3 p r q) *\<^sub>R p + (det3 s p r / det3 p r q) *\<^sub>R q) /\<^sub>R ?sum" unfolding eqs by (simp add: algebra_simps prod_eq_iff) hence srpq: "(det3 p r q / det3 p r q / ?sum) *\<^sub>R s + (det3 q s p / det3 p r q / ?sum) *\<^sub>R r = (det3 r q s / det3 p r q / ?sum) *\<^sub>R p + (det3 s p r / det3 p r q / ?sum) *\<^sub>R q" (is "?s *\<^sub>R s + ?r *\<^sub>R r = ?p *\<^sub>R p + ?q *\<^sub>R q") using det_pos by (simp add: algebra_simps inverse_eq_divide) have eqs: "?s + ?r = 1" "?p + ?q = 1" and s: "?s \ 0" "?s \ 1" and r: "?r \ 0" "?r \ 1" and p: "?p \ 0" "?p \ 1" and q: "?q \ 0" "?q \ 1" unfolding add_divide_distrib[symmetric] using det_pos nonnegs dets_eq by (auto) from eqs have eqs': "1 - ?s = ?r" "1 - ?r = ?s" "1 - ?p = ?q" "1 - ?q = ?p" by auto have comm: "?r *\<^sub>R r + ?s *\<^sub>R s = ?s *\<^sub>R s + ?r *\<^sub>R r" "?q *\<^sub>R q + ?p *\<^sub>R p = ?p *\<^sub>R p + ?q *\<^sub>R q" by simp_all define K where "K = (det3 r q s / det3 p r q / ?sum) *\<^sub>R p + (det3 s p r / det3 p r q / ?sum) *\<^sub>R q" note rewrs = eqs' comm srpq K_def[symmetric] from lex_convex_self2[OF _ s, of s r, unfolded rewrs] lex_convex_self2[OF _ r, of r s, unfolded rewrs] lex_convex_self2[OF _ p, of p q, unfolded rewrs] lex_convex_self2[OF _ q, of q p, unfolded rewrs] have False using g14 d det_pos by (metis lex_trans not_lex_eq) } note wlog = this from dets_eq have 1: "det3 q s p + det3 p r q = det3 s p r + det3 r q s" by simp from d have d': "distinct4 r q p s" by auto note wlog[of q s p r, OF _ 1 nonnegs(3,2,1) prems(4,3,2,1) d'] wlog[of p r q s, OF _ dets_eq nonnegs(2,3,4) prems(1-4) d] ultimately show False using nonnegs d * by (auto simp: less_eq_real_def det3_def' algebra_simps) qed ultimately have ?th1 ?th2 ?th3 by blast+ } ultimately show ?th1 ?th2 ?th3 by force+ qed lemma ccw'_subst_psi_disj: assumes "det3 t r s = 0" assumes "psi t r s \ psi t s r \ psi s r t" assumes "s \ t" assumes "ccw' t r p" shows "ccw' t s p" proof cases assume "r \ s" from assms have "r \ t" by (auto simp: det3_def' ccw'_def algebra_simps) from assms have "det3 r s t = 0" by (auto simp: algebra_simps det3_def') from coll_ex_scaling[OF assms(3) this] obtain x where s: "r = s + x *\<^sub>R (t - s)" by auto from assms(4)[simplified s] have "0 < det3 0 (s + x *\<^sub>R (t - s) - t) (p - t)" by (auto simp: algebra_simps det3_def' ccw'_def) also have "s + x *\<^sub>R (t - s) - t = (1 - x) *\<^sub>R (s - t)" by (simp add: algebra_simps) finally have ccw': "ccw' 0 ((1 - x) *\<^sub>R (s - t)) (p - t)" by (simp add: ccw'_def) hence neq: "x \ 1" by (auto simp add: det3_def' ccw'_def) have tr: "fst s < fst r \ fst t = fst s \ snd t \ snd r" by (simp add: s) from s have "fst (r - s) = fst (x *\<^sub>R (t - s))" "snd (r - s) = snd (x *\<^sub>R (t - s))" by (auto simp: ) hence "x = (if fst (t - s) = 0 then snd (r - s) / snd (t - s) else fst (r - s) / fst (t - s))" using \s \ t\ by (auto simp add: field_simps prod_eq_iff) also have "\ \ 1" using assms by (auto simp: lex_def psi_def tr) finally have "x < 1" using neq by simp thus ?thesis using ccw' by (auto simp: ccw'.translate_origin) qed (insert assms, simp) lemma lex_contr: assumes "distinct4 t s q r" assumes "lex t s" "lex s r" assumes "det3 t s r = 0" assumes "ccw' t s q" assumes "ccw' t q r" shows "False" using ccw'_subst_psi_disj[of t s r q] assms by (cases "r = t") (auto simp: det3_def' algebra_simps psi_def ccw'_def) lemma contra4: assumes "distinct4 s r q p" assumes lex: "lex q p" "lex p r" "lex r s" assumes ccw: "ccw r q s" "ccw r s p" "ccw r q p" shows False proof cases assume c: "ccw s q p" from c have *: "indelta s r q p" using assms by simp with contra1[OF assms(1)] have "\ (lex r q \ lex q p \ lex p s)" by blast hence "\ lex q p" using \ccw s q p\ contra1 cyclic assms nondegenerate by blast thus False using assms by simp next assume "\ ccw s q p" with ccw have "ccw q s p \ ccw s p r \ ccw p r q \ ccw r q s" by (metis assms(1) ccw'.cyclic ccw_def not_ccw'_eq psi_disjuncts) moreover from lex have "lex q r" "lex q s" "lex p r" "lex p s" by order+ ultimately show False using contra3[of r q p s] \distinct4 s r q p\ by blast qed lemma not_coll_ordered_lexI: assumes "l \ x0" and "lex x1 r" and "lex x1 l" and "lex r x0" and "lex l x0" and "ccw' x0 l x1" and "ccw' x0 x1 r" shows "det3 x0 l r \ 0" proof assume "coll x0 l r" from \coll x0 l r\ have 1: "coll 0 (l - x0) (r - x0)" by (simp add: det3_def' algebra_simps) from \lex r x0\ have 2: "lex (r - x0) 0" by (auto simp add: lex_def) from \lex l x0\ have 3: "lex (l - x0) 0" by (auto simp add: lex_def) from \ccw' x0 l x1\ have 4: "ccw' 0 (l - x0) (x1 - x0)" by (simp add: det3_def' ccw'_def algebra_simps) from \ccw' x0 x1 r\ have 5: "ccw' 0 (x1 - x0) (r - x0)" by (simp add: det3_def' ccw'_def algebra_simps) from \lex x1 r\ have 6: "lex 0 (r - x0 - (x1 - x0))" by (auto simp: lex_def) from \lex x1 l\ have 7: "lex 0 (l - x0 - (x1 - x0))" by (auto simp: lex_def) define r' where "r' = r - x0" define l' where "l' = l - x0" define x0' where "x0' = x1 - x0" from 1 2 3 4 5 6 7 have rs: "coll 0 l' r'" "lex r' 0" "lex l' 0" "ccw' 0 l' x0'" "ccw' 0 x0' r'" "lex 0 (r' - x0')" "lex 0 (l' - x0')" unfolding r'_def[symmetric] l'_def[symmetric] x0'_def[symmetric] by auto from assms have "l' \ 0" by (auto simp: l'_def) from coll_scale[OF \coll 0 l' _\ this] obtain y where y: "r' = y *\<^sub>R l'" by auto { assume "y > 0" with rs have False by (auto simp: det3_def' algebra_simps y ccw'_def) } moreover { assume "y < 0" with rs have False - by (auto simp: lex_def not_less algebra_simps sign_simps y ccw'_def) + by (auto simp: lex_def not_less algebra_simps algebra_split_simps y ccw'_def) } moreover { assume "y = 0" from this rs have False by (simp add: ccw'_def y) } ultimately show False by arith qed interpretation ccw_system4 ccw proof unfold_locales fix p q r t assume ccw: "ccw t q r" "ccw p t r" "ccw p q t" show "ccw p q r" proof (cases "det3 t q r = 0 \ det3 p t r = 0 \ det3 p q t = 0") case True { assume "psi t q r \ psi q r t \ psi r t q" "psi p t r \ psi t r p \ psi r p t" "psi p q t \ psi q t p \ psi t p q" hence "psi p q r \ psi q r p \ psi r p q" using lex_sym_eq_iff psi_def by blast } with True ccw show ?thesis by (simp add: det3_def' algebra_simps ccw_def ccw'_def) next case False hence "0 \ det3 t q r" "0 \ det3 p t r" "0 \ det3 p q t" using ccw by (auto simp: less_eq_real_def ccw_def ccw'_def) with False show ?thesis by (auto simp: ccw_def det3_def' algebra_simps ccw'_def intro!: disjI1) qed qed lemma lex_total: "lex t q \ t \ q \ lex q t \ t \ q \ t = q" by auto lemma ccw_two_up_contra: assumes c: "ccw' t p q" "ccw' t q r" assumes ccws: "ccw t s p" "ccw t s q" "ccw t s r" "ccw t p q" "ccw t q r" "ccw t r p" assumes distinct: "distinct5 t s p q r" shows False proof - from ccws have nn: "det3 t s p \ 0" "det3 t s q \ 0" "det3 t s r \ 0" "det3 t r p \ 0" by (auto simp add: less_eq_real_def ccw_def ccw'_def) with c det_identity[of t p q s r] have tsr: "coll t s r" and tsp: "coll t s p" by (auto simp: add_nonneg_eq_0_iff ccw'_def) moreover have trp: "coll t r p" by (metis ccw'_subst_collinear distinct not_ccw'_eq tsr tsp) ultimately have tpr: "coll t p r" by (auto simp: det3_def' algebra_simps) moreover have psi: "psi t p r \ psi t r p \ psi r p t" unfolding psi_def proof - have ntsr: "\ ccw' t s r" "\ ccw' t r s" using tsr by (auto simp: not_ccw'_eq det3_def' algebra_simps) have f8: "\ ccw' t r s" using tsr not_ccw'_eq by blast have f9: "\ ccw' t r p" using tpr by (simp add: not_ccw'_eq) have f10: "(lex t r \ lex r p \ lex r p \ lex p t \ lex p t \ lex t r)" using ccw_def ccws(6) psi_def f9 by auto have "\ ccw' t r q" using c(2) not_ccw'_eq by blast moreover have "\coll t q s" using ntsr ccw'_subst_collinear distinct c(2) by blast hence "ccw' t s q" by (meson ccw_def ccws(2) not_ccw'_eq) moreover from tsr tsp \coll t r p\ have "coll t p s" "coll t p r" "coll t r s" by (auto simp add: det3_def' algebra_simps) ultimately show "lex t p \ lex p r \ lex t r \ lex r p \ lex r p \ lex p t" by (metis ccw'_subst_psi_disj distinct ccw_def ccws(3) contra4 tsp ntsr(1) f10 lex_total psi_def trp) qed moreover from distinct have "r \ t" by auto ultimately have "ccw' t r q" using c(1) by (rule ccw'_subst_psi_disj) thus False using c(2) by (simp add: ccw'_contra) qed lemma ccw_transitive_contr: fixes t s p q r assumes ccws: "ccw t s p" "ccw t s q" "ccw t s r" "ccw t p q" "ccw t q r" "ccw t r p" assumes distinct: "distinct5 t s p q r" shows False proof - from ccws distinct have *: "ccw p t r" "ccw p q t" by (metis cyclic)+ with distinct have "ccw r p q" using interior[OF _ _ ccws(5) *, of UNIV] by (auto intro: cyclic) from ccws have nonnegs: "det3 t s p \ 0" "det3 t s q \ 0" "det3 t s r \ 0" "det3 t p q \ 0" "det3 t q r \ 0" "det3 t r p \ 0" by (auto simp add: less_eq_real_def ccw_def ccw'_def) { assume "ccw' t p q" "ccw' t q r" "ccw' t r p" hence False using ccw_two_up_contra ccws distinct by blast } moreover { assume c: "coll t q r" "coll t r p" with distinct four_points_aligned(1)[OF c, of s] have "coll t p q" by auto hence "(psi t p q \ psi p q t \ psi q t p)" "psi t q r \ psi q r t \ psi r t q" "psi t r p \ psi r p t \ psi p t r" using ccws(4,5,6) c by (simp_all add: ccw_def ccw'_def) hence False using distinct by (auto simp: psi_def ccw'_def) } moreover { assume c: "det3 t p q = 0" "det3 t q r > 0" "det3 t r p = 0" have "\x. det3 t q r = 0 \ t = x \ r = q \ q = x \ r = p \ p = x \ r = x" by (meson c(1) c(3) distinct four_points_aligned(1)) hence False by (metis (full_types) c(2) distinct less_irrefl) } moreover { assume c: "det3 t p q = 0" "det3 t q r = 0" "det3 t r p > 0" have "\x. det3 t r p = 0 \ t = x \ r = x \ q = x \ p = x" by (meson c(1) c(2) distinct four_points_aligned(1)) hence False by (metis (no_types) c(3) distinct less_numeral_extra(3)) } moreover { assume c: "ccw' t p q" "ccw' t q r" from ccw_two_up_contra[OF this ccws distinct] have False . } moreover { assume c: "ccw' t p q" "ccw' t r p" from ccw_two_up_contra[OF this(2,1), of s] ccws distinct have False by auto } moreover { assume c: "ccw' t q r" "ccw' t r p" from ccw_two_up_contra[OF this, of s] ccws distinct have False by auto } ultimately show "False" using \0 \ det3 t p q\ \0 \ det3 t q r\\0 \ det3 t r p\ by (auto simp: less_eq_real_def ccw'_def) qed interpretation ccw: ccw_system ccw by unfold_locales (metis ccw_transitive_contr nondegenerate) lemma ccw_scaleR1: "det3 0 xr P \ 0 \ 0 < e \ ccw 0 xr P \ ccw 0 (e*\<^sub>Rxr) P" by (simp add: ccw_def) lemma ccw_scaleR2: "det3 0 xr P \ 0 \ 0 < e \ ccw 0 xr P \ ccw 0 xr (e*\<^sub>RP)" by (simp add: ccw_def) lemma ccw_translate3_aux: assumes "\coll 0 a b" assumes "x < 1" assumes "ccw 0 (a - x*\<^sub>Ra) (b - x *\<^sub>R a)" shows "ccw 0 a b" proof - from assms have "\ coll 0 (a - x*\<^sub>Ra) (b - x *\<^sub>R a)" by simp with assms have "ccw' 0 ((1 - x) *\<^sub>R a) (b - x *\<^sub>R a)" by (simp add: algebra_simps ccw_def) thus "ccw 0 a b" using \x < 1\ by (simp add: ccw_def) qed lemma ccw_translate3_minus: "det3 0 a b \ 0 \ x < 1 \ ccw 0 a (b - x *\<^sub>R a) \ ccw 0 a b" using ccw_translate3_aux[of a b x] ccw_scaleR1[of a "b - x *\<^sub>R a" "1-x" ] by (auto simp add: algebra_simps) lemma ccw_translate3: "det3 0 a b \ 0 \ x < 1 \ ccw 0 a b \ ccw 0 a (x *\<^sub>R a + b)" by (rule ccw_translate3_minus) (auto simp add: algebra_simps) lemma ccw_switch23: "det3 0 P Q \ 0 \ (\ ccw 0 Q P \ ccw 0 P Q)" by (auto simp: ccw_def algebra_simps not_ccw'_eq ccw'_not_coll) lemma ccw0_upward: "fst a > 0 \ snd a = 0 \ snd b > snd a \ ccw 0 a b" by (auto simp: ccw_def det3_def' algebra_simps ccw'_def) lemma ccw_uminus3[simp]: "det3 a b c \ 0 \ ccw (-a) (-b) (-c) = ccw a b c" by (auto simp: ccw_def ccw'_def algebra_simps det3_def') lemma coll_minus_eq: "coll (x - a) (x - b) (x - c) = coll a b c" by (auto simp: det3_def' algebra_simps) lemma ccw_minus3: "\ coll a b c \ ccw (x - a) (x - b) (x - c) \ ccw a b c" by (simp add: ccw_def coll_minus_eq) lemma ccw0_uminus[simp]: "\ coll 0 a b \ ccw 0 (-a) (-b) \ ccw 0 a b" using ccw_uminus3[of 0 a b] by simp lemma lex_convex2: assumes "lex p q" "lex p r" "0 \ u" "u \ 1" shows "lex p (u *\<^sub>R q + (1 - u) *\<^sub>R r)" proof cases note \lex p q\ also assume "lex q r" hence "lex q (u *\<^sub>R q + (1 - u) *\<^sub>R r)" using \0 \ u\ \u \ 1\ by (rule lex_convex_self2) finally (lex_trans) show ?thesis . next note \lex p r\ also assume "\ lex q r" hence "lex r q" by simp hence "lex r ((1 - u) *\<^sub>R r + (1 - (1 - u)) *\<^sub>R q)" using \0 \ u\ \u \ 1\ by (intro lex_convex_self2) simp_all finally (lex_trans) show ?thesis by (simp add: ac_simps) qed lemma lex_convex2': assumes "lex q p" "lex r p" "0 \ u" "u \ 1" shows "lex (u *\<^sub>R q + (1 - u) *\<^sub>R r) p" proof - have "lex (- p) (u *\<^sub>R (-q) + (1 - u) *\<^sub>R (-r))" using assms by (intro lex_convex2) (auto simp: lex_def) thus ?thesis by (auto simp: lex_def algebra_simps) qed lemma psi_convex1: assumes "psi c a b" assumes "psi d a b" assumes "0 \ u" "0 \ v" "u + v = 1" shows "psi (u *\<^sub>R c + v *\<^sub>R d) a b" proof - from assms have v: "v = (1 - u)" by simp show ?thesis using assms by (auto simp: psi_def v intro!: lex_convex2' lex_convex2) qed lemma psi_convex2: assumes "psi a c b" assumes "psi a d b" assumes "0 \ u" "0 \ v" "u + v = 1" shows "psi a (u *\<^sub>R c + v *\<^sub>R d) b" proof - from assms have v: "v = (1 - u)" by simp show ?thesis using assms by (auto simp: psi_def v intro!: lex_convex2' lex_convex2) qed lemma psi_convex3: assumes "psi a b c" assumes "psi a b d" assumes "0 \ u" "0 \ v" "u + v = 1" shows "psi a b (u *\<^sub>R c + v *\<^sub>R d)" proof - from assms have v: "v = (1 - u)" by simp show ?thesis using assms by (auto simp: psi_def v intro!: lex_convex2) qed lemma coll_convex: assumes "coll a b c" "coll a b d" assumes "0 \ u" "0 \ v" "u + v = 1" shows "coll a b (u *\<^sub>R c + v *\<^sub>R d)" proof cases assume "a \ b" with assms(1, 2) obtain x y where xy: "c - a = x *\<^sub>R (b - a)" "d - a = y *\<^sub>R (b - a)" by (auto simp: det3_translate_origin dest!: coll_scale) from assms have "(u + v) *\<^sub>R a = a" by simp hence "u *\<^sub>R c + v *\<^sub>R d - a = u *\<^sub>R (c - a) + v *\<^sub>R (d - a)" by (simp add: algebra_simps) also have "\ = u *\<^sub>R x *\<^sub>R (b - a) + v *\<^sub>R y *\<^sub>R (b - a)" by (simp add: xy) also have "\ = (u * x + v * y) *\<^sub>R (b - a)" by (simp add: algebra_simps) also have "coll 0 (b - a) \" by (simp add: coll_scaleR_right_eq) finally show ?thesis by (auto simp: det3_translate_origin) qed simp lemma (in ccw_vector_space) convex3: assumes "u \ 0" "v \ 0" "u + v = 1" "ccw a b d" "ccw a b c" shows "ccw a b (u *\<^sub>R c + v *\<^sub>R d)" proof - have "v = 1 - u" using assms by simp hence "ccw 0 (b - a) (u *\<^sub>R (c - a) + v *\<^sub>R (d - a))" using assms by (cases "u = 0" "v = 0" rule: bool.exhaust[case_product bool.exhaust]) (auto simp add: translate_origin intro!: add3) also have "(u + v) *\<^sub>R a = a" by (simp add: assms) hence "u *\<^sub>R (c - a) + v *\<^sub>R (d - a) = u *\<^sub>R c + v *\<^sub>R d - a" by (auto simp: algebra_simps) finally show ?thesis by (simp add: translate_origin) qed lemma ccw_self[simp]: "ccw a a b" "ccw b a a" by (auto simp: ccw_def psi_def intro: cyclic) lemma ccw_sefl'[simp]: "ccw a b a" by (rule cyclic) simp lemma ccw_convex': assumes uv: "u \ 0" "v \ 0" "u + v = 1" assumes "ccw a b c" and 1: "coll a b c" assumes "ccw a b d" and 2: "\ coll a b d" shows "ccw a b (u *\<^sub>R c + v *\<^sub>R d)" proof - from assms have u: "0 \ u" "u \ 1" and v: "v = 1 - u" by (auto simp: algebra_simps) let ?c = "u *\<^sub>R c + v *\<^sub>R d" from 1 have abd: "ccw' a b d" using assms by (auto simp: ccw_def) { assume 2: "\ coll a b c" from 2 have "ccw' a b c" using assms by (auto simp: ccw_def) with abd have "ccw' a b ?c" using assms by (auto intro!: ccw'.convex3) hence ?thesis by (simp add: ccw_def) } moreover { assume 2: "coll a b c" { assume "a = b" hence ?thesis by simp } moreover { assume "v = 0" hence ?thesis by (auto simp: v assms) } moreover { assume "v \ 0" "a \ b" have "coll c a b" using 2 by (auto simp: det3_def' algebra_simps) from coll_ex_scaling[OF \a \ b\ this] obtain r where c: "c = a + r *\<^sub>R (b - a)" by auto have *: "u *\<^sub>R (a + r *\<^sub>R (b - a)) + v *\<^sub>R d - a = (u * r) *\<^sub>R (b - a) + (1 - u) *\<^sub>R (d - a)" by (auto simp: algebra_simps v) have "ccw' a b ?c" using \v \ 0\ uv abd by (simp add: ccw'.translate_origin c *) hence ?thesis by (simp add: ccw_def) } ultimately have ?thesis by blast } ultimately show ?thesis by blast qed lemma ccw_convex: assumes uv: "u \ 0" "v \ 0" "u + v = 1" assumes "ccw a b c" assumes "ccw a b d" assumes lex: "coll a b c \ coll a b d \ lex b a" shows "ccw a b (u *\<^sub>R c + v *\<^sub>R d)" proof - from assms have u: "0 \ u" "u \ 1" and v: "v = 1 - u" by (auto simp: algebra_simps) let ?c = "u *\<^sub>R c + v *\<^sub>R d" { assume coll: "coll a b c \ coll a b d" hence "coll a b ?c" by (auto intro!: coll_convex assms) moreover from coll have "psi a b c \ psi b c a \ psi c a b" "psi a b d \ psi b d a \ psi d a b" using assms by (auto simp add: ccw_def ccw'_not_coll) hence "psi a b ?c \ psi b ?c a \ psi ?c a b" using coll uv lex by (auto simp: psi_def ccw_def not_lex lexs_def v intro: lex_convex2 lex_convex2') ultimately have ?thesis by (simp add: ccw_def) } moreover { assume 1: "\ coll a b d" and 2: "\ coll a b c" from 1 have abd: "ccw' a b d" using assms by (auto simp: ccw_def) from 2 have "ccw' a b c" using assms by (auto simp: ccw_def) with abd have "ccw' a b ?c" using assms by (auto intro!: ccw'.convex3) hence ?thesis by (simp add: ccw_def) } moreover { assume "\ coll a b d" "coll a b c" have ?thesis by (rule ccw_convex') fact+ } moreover { assume 1: "coll a b d" and 2: "\ coll a b c" have "0 \ 1 - u" using assms by (auto ) from ccw_convex'[OF this \0 \ u\ _ \ccw a b d\ 1 \ccw a b c\ 2] have ?thesis by (simp add: algebra_simps v) } ultimately show ?thesis by blast qed interpretation ccw: ccw_convex ccw S "\a b. lex b a" for S by unfold_locales (rule ccw_convex) lemma ccw_sorted_scaleR: "ccw.sortedP 0 xs \ r > 0 \ ccw.sortedP 0 (map ((*\<^sub>R) r) xs)" by (induct xs) (auto intro!: ccw.sortedP.Cons ccw_scale23 elim!: ccw.sortedP_Cons simp del: scaleR_Pair) lemma ccw_sorted_implies_ccw'_sortedP: assumes nonaligned: "\y z. y \ set Ps \ z \ set Ps \ y \ z \ \ coll 0 y z" assumes sorted: "linorder_list0.sortedP (ccw 0) Ps" assumes "distinct Ps" shows "linorder_list0.sortedP (ccw' 0 ) Ps" using assms proof (induction Ps) case (Cons P Ps) { fix p assume p: "p \ set Ps" moreover from p Cons.prems have "ccw 0 P p" by (auto elim!: linorder_list0.sortedP_Cons intro: Cons) ultimately have "ccw' 0 P p" using \distinct (P#Ps)\ by (intro ccw_ncoll_imp_ccw Cons) auto } moreover have "linorder_list0.sortedP (ccw' 0) Ps" using Cons.prems by (intro Cons) (auto elim!: linorder_list0.sortedP_Cons intro: Cons) ultimately show ?case by (auto intro!: linorder_list0.Cons ) qed (auto intro: linorder_list0.Nil) 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/Buchi_Complementation/Complementation_Final.thy b/thys/Buchi_Complementation/Complementation_Final.thy --- a/thys/Buchi_Complementation/Complementation_Final.thy +++ b/thys/Buchi_Complementation/Complementation_Final.thy @@ -1,89 +1,89 @@ section \Complementation to Explicit Büchi Automaton\ theory Complementation_Final imports "Complementation_Implement" "Transition_Systems_and_Automata.NBA_Translate" "HOL-Library.Permutation" begin definition "hci k \ uint32_of_nat k * 1103515245 + 12345" definition "hc \ \ (p, q, b). hci p + hci q * 31 + (if b then 1 else 0)" definition "list_hash xs \ fold (bitXOR \ hc) xs 0" lemma list_hash_eq: assumes "distinct xs" "distinct ys" "set xs = set ys" shows "list_hash xs = list_hash ys" proof - have "remdups xs <~~> remdups ys" using eq_set_perm_remdups assms(3) by this then have "xs <~~> ys" using assms(1, 2) by (simp add: distinct_remdups_id) then have "fold (bitXOR \ hc) xs a = fold (bitXOR \ hc) ys a" for a proof (induct arbitrary: a) case (swap y x l) have "x XOR y XOR a = y XOR x XOR a" for x y by (transfer) (simp add: word_bw_lcs(3)) then show ?case by simp qed simp+ then show ?thesis unfolding list_hash_def by this qed definition state_hash :: "nat \ Complementation_Implement.state \ nat" where "state_hash n p \ nat_of_hashcode (list_hash p) mod n" lemma state_hash_bounded_hashcode[autoref_ga_rules]: "is_bounded_hashcode state_rel (gen_equals (Gen_Map.gen_ball (foldli \ list_map_to_list)) (list_map_lookup (=)) (prod_eq (=) (\))) state_hash" proof show [param]: "(gen_equals (Gen_Map.gen_ball (foldli \ list_map_to_list)) (list_map_lookup (=)) (prod_eq (=) (\)), (=)) \ state_rel \ state_rel \ bool_rel" by autoref show "state_hash n xs = state_hash n ys" if "xs \ Domain state_rel" "ys \ Domain state_rel" "gen_equals (Gen_Map.gen_ball (foldli \ list_map_to_list)) (list_map_lookup (=)) (prod_eq (=) (=)) xs ys" for xs ys n proof - have 1: "distinct (map fst xs)" "distinct (map fst ys)" using that(1, 2) unfolding list_map_rel_def list_map_invar_def by (auto simp: in_br_conv) have 2: "distinct xs" "distinct ys" using 1 by (auto intro: distinct_mapI) have 3: "(xs, map_of xs) \ state_rel" "(ys, map_of ys) \ state_rel" using 1 unfolding list_map_rel_def list_map_invar_def by (auto simp: in_br_conv) have 4: "(gen_equals (Gen_Map.gen_ball (foldli \ list_map_to_list)) (list_map_lookup (=)) (prod_eq (=) (\)) xs ys, map_of xs = map_of ys) \ bool_rel" using 3 by parametricity have 5: "map_to_set (map_of xs) = map_to_set (map_of ys)" using that(3) 4 by simp have 6: "set xs = set ys" using map_to_set_map_of 1 5 by blast show "state_hash n xs = state_hash n ys" unfolding state_hash_def using list_hash_eq 2 6 by metis qed show "state_hash n x < n" if "1 < n" for n x using that unfolding state_hash_def by simp qed schematic_goal complement_impl: assumes [simp]: "finite (nodes A)" assumes [autoref_rules]: "(Ai, A) \ \Id, nat_rel\ nbai_nba_rel" shows "(?f :: ?'c, RETURN (to_nbaei (complement_4 A))) \ ?R" by (autoref_monadic (plain)) concrete_definition complement_impl uses complement_impl[unfolded autoref_tag_defs] theorem complement_impl_correct: assumes "finite (nodes A)" assumes "(Ai, A) \ \Id, nat_rel\ nbai_nba_rel" shows "language (nbae_nba (nbaei_nbae (complement_impl Ai))) = streams (alphabet A) - language A" proof - have "(language ((nbae_nba \ nbaei_nbae) (complement_impl Ai)), language (id (complement_4 A))) \ \\Id_on (alphabet (complement_4 A))\ stream_rel\ set_rel" using complement_impl.refine[OF assms, unfolded to_nbaei_def id_apply, THEN RETURN_nres_relD] by parametricity also have "language (id (complement_4 A)) = streams (alphabet A) - language A" using assms(1) complement_4_correct by simp finally show ?thesis by simp qed definition nbaei_nbai :: "(String.literal, nat) nbaei \ (String.literal, nat) nbai" where - "nbaei_nbai \ nbae_nba_impl" + "nbaei_nbai \ nbae_nba_impl (=) (=)" (* TODO: possible optimizations: - introduce op_map_map operation for maps instead of manually iterating via FOREACH - consolidate various binds and maps in expand_map_get_7 *) export_code nat_of_integer integer_of_nat nbaei alphabetei initialei transitionei acceptingei nbaei_nbai complement_impl in SML module_name Complementation file_prefix Complementation_Export end \ No newline at end of file diff --git a/thys/Fourier/Fourier.thy b/thys/Fourier/Fourier.thy --- a/thys/Fourier/Fourier.thy +++ b/thys/Fourier/Fourier.thy @@ -1,2742 +1,2742 @@ section\The basics of Fourier series\ text\Ported from HOL Light; thanks to Manuel Eberl for help with the real asymp proof methods\ theory "Fourier" imports Periodic Square_Integrable "HOL-Real_Asymp.Real_Asymp" Confine Fourier_Aux2 begin subsection\Orthonormal system of L2 functions and their Fourier coefficients\ definition orthonormal_system :: "'a::euclidean_space set \ ('b \ 'a \ real) \ bool" where "orthonormal_system S w \ \m n. l2product S (w m) (w n) = (if m = n then 1 else 0)" definition orthonormal_coeff :: "'a::euclidean_space set \ (nat \ 'a \ real) \ ('a \ real) \ nat \ real" where "orthonormal_coeff S w f n = l2product S (w n) f" lemma orthonormal_system_eq: "orthonormal_system S w \ l2product S (w m) (w n) = (if m = n then 1 else 0)" by (simp add: orthonormal_system_def) lemma orthonormal_system_l2norm: "orthonormal_system S w \ l2norm S (w i) = 1" by (simp add: l2norm_def orthonormal_system_def) lemma orthonormal_partial_sum_diff: assumes os: "orthonormal_system S w" and w: "\i. (w i) square_integrable S" and f: "f square_integrable S" and "finite I" shows "(l2norm S (\x. f x - (\i\I. a i * w i x)))\<^sup>2 = (l2norm S f)\<^sup>2 + (\i\I. (a i)\<^sup>2) - 2 * (\i\I. a i * orthonormal_coeff S w f i)" proof - have "S \ sets lebesgue" using f square_integrable_def by blast then have "(\x. \i\I. a i * w i x) square_integrable S" by (intro square_integrable_sum square_integrable_lmult w \finite I\) with assms show ?thesis - apply (simp add: square_integrable_diff l2norm_pow_2 orthonormal_coeff_def orthonormal_system_def) - apply (simp add: square_integrable_diff square_integrable_lmult l2product_rdiff l2product_sym - l2product_rsum l2product_rmult algebra_simps power2_eq_square if_distrib [of "\x. _ * x"]) + apply (simp add: l2norm_pow_2 orthonormal_coeff_def orthonormal_system_def) + apply (simp add: l2product_rdiff l2product_sym + l2product_rsum l2product_rmult algebra_simps power2_eq_square if_distrib [of "\x. _ * x"]) done qed lemma orthonormal_optimal_partial_sum: assumes "orthonormal_system S w" "\i. (w i) square_integrable S" "f square_integrable S" "finite I" shows "l2norm S (\x. f x - (\i\I. orthonormal_coeff S w f i * w i x)) \ l2norm S (\x. f x - (\i\I. a i * w i x))" proof - have "2 * (a i * l2product S f(w i)) \ (a i)\<^sup>2 + (l2product S f(w i))\<^sup>2" for i using sum_squares_bound [of "a i" "l2product S f(w i)"] by (force simp: algebra_simps) then have *: "2 * (\i\I. a i * l2product S f(w i)) \ (\i\I. (a i)\<^sup>2 + (l2product S f(w i))\<^sup>2)" by (simp add: sum_distrib_left sum_mono) have S: "S \ sets lebesgue" using assms square_integrable_def by blast with assms * show ?thesis apply (simp add: l2norm_le square_integrable_sum square_integrable_diff square_integrable_lmult flip: l2norm_pow_2) apply (simp add: orthonormal_coeff_def orthonormal_partial_sum_diff) apply (simp add: l2product_sym power2_eq_square sum.distrib) done qed lemma Bessel_inequality: assumes "orthonormal_system S w" "\i. (w i) square_integrable S" "f square_integrable S" "finite I" shows "(\i\I. (orthonormal_coeff S w f i)\<^sup>2) \ (l2norm S f)\<^sup>2" using orthonormal_partial_sum_diff [OF assms, of "orthonormal_coeff S w f"] apply (simp add: algebra_simps flip: power2_eq_square) by (metis (lifting) add_diff_cancel_left' diff_ge_0_iff_ge zero_le_power2) lemma Fourier_series_square_summable: assumes os: "orthonormal_system S w" and w: "\i. (w i) square_integrable S" and f: "f square_integrable S" shows "summable (confine (\i. (orthonormal_coeff S w f i) ^ 2) I)" proof - have "incseq (\n. \i\I \ {..n}. (orthonormal_coeff S w f i)\<^sup>2)" by (auto simp: incseq_def intro: sum_mono2) moreover have "\i. (\i\I \ {..i}. (orthonormal_coeff S w f i)\<^sup>2) \ (l2norm S f)\<^sup>2" by (simp add: Bessel_inequality assms) ultimately obtain L where "(\n. \i\I \ {..n}. (orthonormal_coeff S w f i)\<^sup>2) \ L" using incseq_convergent by blast then have "\r>0. \no. \n\no. \(\i\{..n} \ I. (orthonormal_coeff S w f i)\<^sup>2) - L\ < r" by (simp add: LIMSEQ_iff Int_commute) then show ?thesis by (auto simp: summable_def sums_confine_le LIMSEQ_iff) qed lemma orthonormal_Fourier_partial_sum_diff_squared: assumes os: "orthonormal_system S w" and w: "\i. (w i) square_integrable S" and f: "f square_integrable S" and "finite I" shows "(l2norm S (\x. f x -(\i\I. orthonormal_coeff S w f i * w i x)))\<^sup>2 = (l2norm S f)\<^sup>2 - (\i\I. (orthonormal_coeff S w f i)\<^sup>2)" using orthonormal_partial_sum_diff [OF assms, where a = "orthonormal_coeff S w f"] by (simp add: power2_eq_square) lemma Fourier_series_l2_summable: assumes os: "orthonormal_system S w" and w: "\i. (w i) square_integrable S" and f: "f square_integrable S" obtains g where "g square_integrable S" "(\n. l2norm S (\x. (\i\I \ {..n}. orthonormal_coeff S w f i * w i x) - g x)) \ 0" proof - have S: "S \ sets lebesgue" using f square_integrable_def by blast let ?\ = "\A x. (\i\I \ A. orthonormal_coeff S w f i * w i x)" let ?\ = "confine (\i. (orthonormal_coeff S w f i)\<^sup>2) I" have si: "?\ A square_integrable S" if "finite A" for A by (force intro: square_integrable_lmult w square_integrable_sum S that) have "\N. \m\N. \n\N. l2norm S (\x. ?\ {..m} x - ?\ {..n} x) < e" if "e > 0" for e proof - have "summable ?\" by (rule Fourier_series_square_summable [OF os w f]) then have "Cauchy (\n. sum ?\ {..n})" by (simp add: summable_def sums_def_le convergent_eq_Cauchy) then obtain M where M: "\m n. \m\M; n\M\ \ dist (sum ?\ {..m}) (sum ?\ {..n}) < e\<^sup>2" using \e > 0\ unfolding Cauchy_def by (drule_tac x="e^2" in spec) auto have "\m\M; n\M\ \ l2norm S (\x. ?\ {..m} x - ?\ {..n} x) < e" for m n proof (induct m n rule: linorder_class.linorder_wlog) case (le m n) have sum_diff: "sum f(I \ {..n}) - sum f(I \ {..m}) = sum f(I \ {Suc m..n})" for f :: "nat \ real" proof - have "(I \ {..n}) = (I \ {..m} \ I \ {Suc m..n})" "(I \ {..m}) \ (I \ {Suc m..n}) = {}" using le by auto then show ?thesis by (simp add: algebra_simps flip: sum.union_disjoint) qed have "(l2norm S (\x. ?\ {..n} x - ?\ {..m} x))^2 \ \(\i\I \ {..n}. (orthonormal_coeff S w f i)\<^sup>2) - (\i\I \ {..m}. (orthonormal_coeff S w f i)\<^sup>2)\" proof (simp add: sum_diff) have "(l2norm S (?\ {Suc m..n}))\<^sup>2 = (\i\I \ {Suc m..n}. l2product S (\x. \j\I \ {Suc m..n}. orthonormal_coeff S w f j * w j x) (\x. orthonormal_coeff S w f i * w i x))" by (simp add: l2norm_pow_2 l2product_rsum si w) also have "\ = (\i\I \ {Suc m..n}. \j\I \ {Suc m..n}. orthonormal_coeff S w f j * orthonormal_coeff S w f i * l2product S (w j) (w i))" by (simp add: l2product_lsum l2product_lmult l2product_rmult si w flip: mult.assoc) also have "\ \ \\i\I \ {Suc m..n}. (orthonormal_coeff S w f i)\<^sup>2\" by (simp add: orthonormal_system_eq [OF os] power2_eq_square if_distrib [of "(*)_"] cong: if_cong) finally show "(l2norm S (?\ {Suc m..n}))\<^sup>2 \ \\i\I \ {Suc m..n}. (orthonormal_coeff S w f i)\<^sup>2\" . qed also have "\ < e\<^sup>2" using M [OF \n\M\ \m\M\] by (simp add: dist_real_def sum_confine_eq_Int Int_commute) finally have "(l2norm S (\x. ?\ {..m} x - ?\ {..n} x))^2 < e^2" using l2norm_diff [OF si si] by simp with \e > 0\ show ?case by (simp add: power2_less_imp_less) next case (sym a b) then show ?case by (subst l2norm_diff) (auto simp: si) qed then show ?thesis by blast qed with that show thesis by (blast intro: si [of "{.._}"] l2_complete [of "\n. ?\ {..n}"]) qed lemma Fourier_series_l2_summable_strong: assumes os: "orthonormal_system S w" and w: "\i. (w i) square_integrable S" and f: "f square_integrable S" obtains g where "g square_integrable S" "\i. i \ I \ orthonormal_coeff S w (\x. f x - g x) i = 0" "(\n. l2norm S (\x. (\i\I \ {..n}. orthonormal_coeff S w f i * w i x) - g x)) \ 0" proof - have S: "S \ sets lebesgue" using f square_integrable_def by blast obtain g where g: "g square_integrable S" and teng: "(\n. l2norm S (\x. (\i\I \ {..n}. orthonormal_coeff S w f i * w i x) - g x)) \ 0" using Fourier_series_l2_summable [OF assms] . show thesis proof show "orthonormal_coeff S w (\x. f x - g x) i = 0" if "i \ I" for i proof (rule tendsto_unique) let ?\ = "\A x. (\i\I \ A. orthonormal_coeff S w f i * w i x)" let ?f = "\n. l2product S (w i) (\x. (f x - ?\ {..n} x) + (?\ {..n} x - g x))" show "?f \ orthonormal_coeff S w (\x. f x - g x) i" by (simp add: orthonormal_coeff_def) have 1: "(\n. l2product S (w i) (\x. f x - ?\ {..n} x)) \ 0" proof (rule tendsto_eventually) have "l2product S (w i) (\x. f x - ?\ {..j} x) = 0" if "j \ i" for j using that \i \ I\ apply (simp add: l2product_rdiff l2product_rsum l2product_rmult orthonormal_coeff_def w f S) apply (simp add: orthonormal_system_eq [OF os] if_distrib [of "(*)_"] cong: if_cong) done then show "\\<^sub>F n in sequentially. l2product S (w i) (\x. f x - ?\ {..n} x) = 0" using eventually_at_top_linorder by blast qed have 2: "(\n. l2product S (w i) (\x. ?\ {..n} x - g x)) \ 0" proof (intro Lim_null_comparison [OF _ teng] eventuallyI) show "norm (l2product S (w i) (\x. ?\ {..n} x - g x)) \ l2norm S (\x. ?\ {..n} x - g x)" for n using Schwartz_inequality_abs [of "w i" S "(\x. ?\ {..n} x - g x)"] by (simp add: S g f w orthonormal_system_l2norm [OF os]) qed show "?f \ 0" using that tendsto_add [OF 1 2] by (subst l2product_radd) (simp_all add: l2product_radd w f g S) qed auto qed (auto simp: g teng) qed subsection\Actual trigonometric orthogonality relations\ lemma integrable_sin_cx: "integrable (lebesgue_on {-pi..pi}) (\x. sin(x * c))" by (intro continuous_imp_integrable_real continuous_intros) lemma integrable_cos_cx: "integrable (lebesgue_on {-pi..pi}) (\x. cos(x * c))" by (intro continuous_imp_integrable_real continuous_intros) lemma integral_cos_Z' [simp]: assumes "n \ \" shows "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(n * x)) = (if n = 0 then 2 * pi else 0)" by (metis assms integral_cos_Z mult_commute_abs) lemma integral_sin_and_cos: fixes m n::int shows "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(m * x) * cos(n * x)) = (if \m\ = \n\ then if n = 0 then 2 * pi else pi else 0)" "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(m * x) * sin(n * x)) = 0" "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin(m * x) * cos(n * x)) = 0" "\m \ 0; n \ 0\ \ integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin (m * x) * sin (n * x)) = (if m = n \ n \ 0 then pi else 0)" "\integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin (m * x) * sin (n * x))\ = (if \m\ = \n\ \ n \ 0 then pi else 0)" by (simp_all add: abs_if sin_times_sin cos_times_sin sin_times_cos cos_times_cos integrable_sin_cx integrable_cos_cx mult_ac flip: distrib_left distrib_right left_diff_distrib right_diff_distrib) lemma integral_sin_and_cos_Z [simp]: fixes m n::real assumes "m \ \" "n \ \" shows "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(m * x) * cos(n * x)) = (if \m\ = \n\ then if n = 0 then 2 * pi else pi else 0)" "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(m * x) * sin(n * x)) = 0" "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin(m * x) * cos(n * x)) = 0" "\integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin (m * x) * sin (n * x))\ = (if \m\ = \n\ \ n \ 0 then pi else 0)" using assms unfolding Ints_def apply safe unfolding integral_sin_and_cos apply auto done lemma integral_sin_and_cos_N [simp]: fixes m n::real assumes "m \ \" "n \ \" shows "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin (m * x) * sin (n * x)) = (if m = n \ n \ 0 then pi else 0)" using assms unfolding Nats_altdef1 by (auto simp: integral_sin_and_cos) lemma integrable_sin_and_cos: fixes m n::int shows "integrable (lebesgue_on {a..b}) (\x. cos(x * m) * cos(x * n))" "integrable (lebesgue_on {a..b}) (\x. cos(x * m) * sin(x * n))" "integrable (lebesgue_on {a..b}) (\x. sin(x * m) * cos(x * n))" "integrable (lebesgue_on {a..b}) (\x. sin(x * m) * sin(x * n))" by (intro continuous_imp_integrable_real continuous_intros)+ lemma sqrt_pi_ge1: "sqrt pi \ 1" using pi_gt3 by auto definition trigonometric_set :: "nat \ real \ real" where "trigonometric_set n \ if n = 0 then \x. 1 / sqrt(2 * pi) else if odd n then \x. sin(real(Suc (n div 2)) * x) / sqrt(pi) else (\x. cos((n div 2) * x) / sqrt pi)" lemma trigonometric_set: "trigonometric_set 0 x = 1 / sqrt(2 * pi)" "trigonometric_set (Suc (2 * n)) x = sin(real(Suc n) * x) / sqrt(pi)" "trigonometric_set (2 * n + 2) x = cos(real(Suc n) * x) / sqrt(pi)" "trigonometric_set (Suc (Suc (2 * n))) x = cos(real(Suc n) * x) / sqrt(pi)" by (simp_all add: trigonometric_set_def algebra_simps add_divide_distrib) lemma trigonometric_set_even: "trigonometric_set(2*k) = (if k = 0 then (\x. 1 / sqrt(2 * pi)) else (\x. cos(k * x) / sqrt pi))" by (induction k) (auto simp: trigonometric_set_def add_divide_distrib split: if_split_asm) lemma orthonormal_system_trigonometric_set: "orthonormal_system {-pi..pi} trigonometric_set" proof - have "l2product {-pi..pi} (trigonometric_set m) (trigonometric_set n) = (if m = n then 1 else 0)" for m n proof (induction m rule: odd_even_cases) case 0 show ?case - by (induction n rule: odd_even_cases) (auto simp: trigonometric_set l2product_def) + by (induction n rule: odd_even_cases) (auto simp: trigonometric_set l2product_def measure_restrict_space) next case (odd m) show ?case by (induction n rule: odd_even_cases) (auto simp: trigonometric_set l2product_def double_not_eq_Suc_double) next case (even m) show ?case by (induction n rule: odd_even_cases) (auto simp: trigonometric_set l2product_def Suc_double_not_eq_double) qed then show ?thesis unfolding orthonormal_system_def by auto qed lemma square_integrable_trigonometric_set: "(trigonometric_set i) square_integrable {-pi..pi}" proof - have "continuous_on {-pi..pi} (\x. sin ((1 + real n) * x) / sqrt pi)" for n by (intro continuous_intros) auto moreover have "continuous_on {-pi..pi} (\x. cos (real i * x / 2) / sqrt pi) " by (intro continuous_intros) auto ultimately show ?thesis by (simp add: trigonometric_set_def) qed subsection\Weierstrass for trigonometric polynomials\ lemma Weierstrass_trig_1: fixes g :: "real \ real" assumes contf: "continuous_on UNIV g" and periodic: "\x. g(x + 2 * pi) = g x" and 1: "norm z = 1" shows "continuous (at z within (sphere 0 1)) (g \ Im \ Ln)" proof (cases "Re z < 0") let ?f = "complex_of_real \ g \ (\x. x + pi) \ Im \ Ln \ uminus" case True have "continuous (at z within (sphere 0 1)) (complex_of_real \ g \ Im \ Ln)" proof (rule continuous_transform_within) have [simp]: "z \ \\<^sub>\\<^sub>0" using True complex_nonneg_Reals_iff by auto show "continuous (at z within (sphere 0 1)) ?f" by (intro continuous_within_Ln continuous_intros continuous_on_imp_continuous_within [OF contf]) auto show "?f x' = (complex_of_real \ g \ Im \ Ln) x'" if "x' \ (sphere 0 1)" and "dist x' z < 1" for x' proof - have "x' \ 0" using that by auto with that show ?thesis by (auto simp: Ln_minus add.commute periodic) qed qed (use 1 in auto) then have "continuous (at z within (sphere 0 1)) (Re \ complex_of_real \ g \ Im \ Ln)" unfolding o_def by (rule continuous_Re) then show ?thesis by (simp add: o_def) next case False let ?f = "complex_of_real \ g \ Im \ Ln \ uminus" have "z \ 0" using 1 by auto with False have "z \ \\<^sub>\\<^sub>0" by (auto simp: not_less nonpos_Reals_def) then have "continuous (at z within (sphere 0 1)) (complex_of_real \ g \ Im \ Ln)" by (intro continuous_within_Ln continuous_intros continuous_on_imp_continuous_within [OF contf]) auto then have "continuous (at z within (sphere 0 1)) (Re \ complex_of_real \ g \ Im \ Ln)" unfolding o_def by (rule continuous_Re) then show ?thesis by (simp add: o_def) qed inductive_set cx_poly :: "(complex \ real) set" where Re: "Re \ cx_poly" | Im: "Im \ cx_poly" | const: "(\x. c) \ cx_poly" | add: "\f \ cx_poly; g \ cx_poly\ \ (\x. f x + g x) \ cx_poly" | mult: "\f \ cx_poly; g \ cx_poly\ \ (\x. f x * g x) \ cx_poly" declare cx_poly.intros [intro] lemma Weierstrass_trig_polynomial: assumes contf: "continuous_on {-pi..pi} f" and fpi: "f(-pi) = f pi" and "0 < e" obtains n::nat and a b where "\x::real. x \ {-pi..pi} \ \f x - (\k\n. a k * sin (k * x) + b k * cos (k * x))\ < e" proof - interpret CR: function_ring_on "cx_poly" "sphere 0 1" proof show "continuous_on (sphere 0 1) f" if "f \ cx_poly" for f using that by induction (assumption | intro continuous_intros)+ fix x y::complex assume "x \ sphere 0 1" and "y \ sphere 0 1" and "x \ y" then consider "Re x \ Re y" | "Im x \ Im y" using complex_eqI by blast then show "\f\cx_poly. f x \ f y" by (metis cx_poly.Im cx_poly.Re) qed (auto simp: cx_poly.intros) have "continuous (at z within (sphere 0 1)) (f \ Im \ Ln)" if "norm z = 1" for z proof - obtain g where contg: "continuous_on UNIV g" and gf: "\x. x \ {-pi..pi} \ g x = f x" and periodic: "\x. g(x + 2*pi) = g x" using Tietze_periodic_interval [OF contf fpi] by auto let ?f = "(g \ Im \ Ln)" show ?thesis proof (rule continuous_transform_within) show "continuous (at z within (sphere 0 1)) ?f" using Weierstrass_trig_1 [OF contg periodic that] by (simp add: sphere_def) show "?f x' = (f \ Im \ Ln) x'" if "x' \ sphere 0 1" "dist x' z < 1" for x' proof - have "x' \ 0" using that by auto then have "Im (Ln x') \ {-pi..pi}" using Im_Ln_le_pi [of x'] mpi_less_Im_Ln [of x'] by simp then show ?thesis using gf by simp qed qed (use that in auto) qed then have "continuous_on (sphere 0 1) (f \ Im \ Ln)" using continuous_on_eq_continuous_within mem_sphere_0 by blast then obtain g where "g \ cx_poly" and g: "\x. x \ sphere 0 1 \ \(f \ Im \ Ln) x - g x\ < e" using CR.Stone_Weierstrass_basic [of "f \ Im \ Ln"] \e > 0\ by meson define trigpoly where "trigpoly \ \f. \n a b. f = (\x. (\k\n. a k * sin(real k * x) + b k * cos(real k * x)))" have tp_const: "trigpoly(\x. c)" for c unfolding trigpoly_def by (rule_tac x=0 in exI) auto have tp_add: "trigpoly(\x. f x + g x)" if "trigpoly f" "trigpoly g" for f g proof - obtain n1 a1 b1 where feq: "f = (\x. \k\n1. a1 k * sin (real k * x) + b1 k * cos (real k * x))" using \trigpoly f\ trigpoly_def by blast obtain n2 a2 b2 where geq: "g = (\x. \k\n2. a2 k * sin (real k * x) + b2 k * cos (real k * x))" using \trigpoly g\ trigpoly_def by blast let ?a = "\n. (if n \ n1 then a1 n else 0) + (if n \ n2 then a2 n else 0)" let ?b = "\n. (if n \ n1 then b1 n else 0) + (if n \ n2 then b2 n else 0)" have eq: "{k. k \ max a b \ k \ a} = {..a}" "{k. k \ max a b \ k \ b} = {..b}" for a b::nat by auto have "(\x. f x + g x) = (\x. \k \ max n1 n2. ?a k * sin (real k * x) + ?b k * cos (real k * x))" by (simp add: feq geq algebra_simps eq sum.distrib if_distrib [of "\u. _ * u"] cong: if_cong flip: sum.inter_filter) then show ?thesis unfolding trigpoly_def by meson qed have tp_sum: "trigpoly(\x. \i\S. f i x)" if "finite S" "\i. i \ S \ trigpoly(f i)" for f and S :: "nat set" using that by induction (auto simp: tp_const tp_add) have tp_cmul: "trigpoly(\x. c * f x)" if "trigpoly f" for f c proof - obtain n a b where feq: "f = (\x. \k\n. a k * sin (real k * x) + b k * cos (real k * x))" using \trigpoly f\ trigpoly_def by blast have "(\x. c * f x) = (\x. \k \ n. (c * a k) * sin (real k * x) + (c * b k) * cos (real k * x))" by (simp add: feq algebra_simps sum_distrib_left) then show ?thesis unfolding trigpoly_def by meson qed have tp_cdiv: "trigpoly(\x. f x / c)" if "trigpoly f" for f c using tp_cmul [OF \trigpoly f\, of "inverse c"] by (simp add: divide_inverse_commute) have tp_diff: "trigpoly(\x. f x - g x)" if "trigpoly f" "trigpoly g" for f g using tp_add [OF \trigpoly f\ tp_cmul [OF \trigpoly g\, of "-1"]] by auto have tp_sin: "trigpoly(\x. sin (n * x))" if "n \ \" for n unfolding trigpoly_def proof obtain k where "n = real k" using Nats_cases \n \ \\ by blast then show "\a b. (\x. sin (n * x)) = (\x. \i \ nat\n\. a i * sin (real i * x) + b i * cos (real i * x))" apply (rule_tac x="\i. if i = k then 1 else 0" in exI) apply (rule_tac x="\i. 0" in exI) apply (simp add: if_distrib [of "\u. u * _"] cong: if_cong) done qed have tp_cos: "trigpoly(\x. cos (n * x))" if "n \ \" for n unfolding trigpoly_def proof obtain k where "n = real k" using Nats_cases \n \ \\ by blast then show "\a b. (\x. cos (n * x)) = (\x. \i \ nat\n\. a i * sin (real i * x) + b i * cos (real i * x))" apply (rule_tac x="\i. 0" in exI) apply (rule_tac x="\i. if i = k then 1 else 0" in exI) apply (simp add: if_distrib [of "\u. u * _"] cong: if_cong) done qed have tp_sincos: "trigpoly(\x. sin(i * x) * sin(j * x)) \ trigpoly(\x. sin(i * x) * cos(j * x)) \ trigpoly(\x. cos(i * x) * sin(j * x)) \ trigpoly(\x. cos(i * x) * cos(j * x))" (is "?P i j") for i j::nat proof (rule linorder_wlog [of _ j i]) show "?P j i" if "i \ j" for j i using that - by (auto simp add: sin_times_sin cos_times_cos sin_times_cos cos_times_sin diff_divide_distrib - simp flip: of_nat_diff left_diff_distrib distrib_right - intro!: tp_add tp_diff tp_cdiv tp_cos tp_sin) + by (simp add: sin_times_sin cos_times_cos sin_times_cos cos_times_sin diff_divide_distrib + tp_add tp_diff tp_cdiv tp_cos tp_sin flip: left_diff_distrib distrib_right) qed (simp add: mult_ac) have tp_mult: "trigpoly(\x. f x * g x)" if "trigpoly f" "trigpoly g" for f g proof - obtain n1 a1 b1 where feq: "f = (\x. \k\n1. a1 k * sin (real k * x) + b1 k * cos (real k * x))" using \trigpoly f\ trigpoly_def by blast obtain n2 a2 b2 where geq: "g = (\x. \k\n2. a2 k * sin (real k * x) + b2 k * cos (real k * x))" using \trigpoly g\ trigpoly_def by blast then show ?thesis unfolding feq geq by (simp add: algebra_simps sum_product tp_sum tp_add tp_cmul tp_sincos del: mult.commute) qed have *: "trigpoly(\x. f(exp(\ * of_real x)))" if "f \ cx_poly" for f using that proof induction case Re then show ?case using tp_cos [of 1] by (auto simp: Re_exp) next case Im then show ?case using tp_sin [of 1] by (auto simp: Im_exp) qed (auto simp: tp_const tp_add tp_mult) obtain n a b where eq: "(g (iexp x)) = (\k\n. a k * sin (real k * x) + b k * cos (real k * x))" for x using * [OF \g \ cx_poly\] trigpoly_def by meson show thesis proof show "\f \ - (\k\n. a k * sin (real k * \) + b k * cos (real k * \))\ < e" if "\ \ {-pi..pi}" for \ proof - have "f \ - g (iexp \) = (f \ Im \ Ln) (iexp \) - g (iexp \)" proof (cases "\ = -pi") case True then show ?thesis by (simp add: exp_minus fpi) next case False then show ?thesis using that by auto qed then show ?thesis using g [of "exp(\ * of_real \)"] by (simp flip: eq) qed qed qed subsection\A bit of extra hacking round so that the ends of a function are OK\ lemma integral_tweak_ends: fixes a b :: real assumes "a < b" "e > 0" obtains f where "continuous_on {a..b} f" "f a = d" "f b = 0" "l2norm {a..b} f < e" proof - have cont: "continuous_on {a..b} (\x. if x \ a+1 / real(Suc n) then ((Suc n) * d) * ((a+1 / (Suc n)) - x) else 0)" for n proof (cases "a+1/(Suc n) \ b") case True have *: "1 / (1 + real n) > 0" by auto have abeq: "{a..b} = {a..a+1/(Suc n)} \ {a+1/(Suc n)..b}" using \a < b\ True apply auto using "*" by linarith show ?thesis unfolding abeq proof (rule continuous_on_cases) show "continuous_on {a..a+1 / real (Suc n)} (\x. real (Suc n) * d * (a+1 / real (Suc n) - x))" by (intro continuous_intros) qed auto next case False show ?thesis proof (rule continuous_on_eq [where f = "\x. ((Suc n) * d) * ((a+1/(Suc n)) - x)"]) show " continuous_on {a..b} (\x. (Suc n) * d * (a+1 / real (Suc n) - x))" by (intro continuous_intros) qed (use False in auto) qed let ?f = "\k x. (if x \ a+1 / (Suc k) then (Suc k) * d * (a+1 / (Suc k) - x) else 0)\<^sup>2" let ?g = "\x. if x = a then d\<^sup>2 else 0" have bmg: "?g \ borel_measurable (lebesgue_on {a..b})" apply (rule measurable_restrict_space1) using borel_measurable_if_I [of _ "{a}", OF borel_measurable_const] by auto have bmf: "?f k \ borel_measurable (lebesgue_on {a..b})" for k proof - have bm: "(\x. (Suc k) * d * (a+1 / real (Suc k) - x)) \ borel_measurable (lebesgue_on {..a+1 / (Suc k)})" by (intro measurable) (auto simp: measurable_completion measurable_restrict_space1) show ?thesis apply (intro borel_measurable_power measurable_restrict_space1) using borel_measurable_if_I [of _ "{.. a+1 / (Suc k)}", OF bm] apply auto done qed have int_d2: "integrable (lebesgue_on {a..b}) (\x. d\<^sup>2)" by (intro continuous_imp_integrable_real continuous_intros) have "(\k. ?f k x) \ ?g x" if x: "x \ {a..b}" for x proof (cases "x = a") case False then have "x > a" using x by auto with x obtain N where "N > 0" and N: "1 / real N < x-a" using real_arch_invD [of "x-a"] by (force simp: inverse_eq_divide) then have "x > a+1 / (1 + real k)" if "k \ N" for k proof - have "a+1 / (1 + real k) < a+1 / real N" using that \0 < N\ by (simp add: field_simps) also have "\ < x" using N by linarith finally show ?thesis . qed then show ?thesis apply (intro tendsto_eventually eventually_sequentiallyI [where c=N]) by (fastforce simp: False) qed auto then have tends: "AE x in (lebesgue_on {a..b}). (\k. ?f k x) \ ?g x" by force have le_d2: "\k. AE x in (lebesgue_on {a..b}). norm (?f k x) \ d\<^sup>2" proof show "norm ((if x \ a+1 / real (Suc k) then real (Suc k) * d * (a+1 / real (Suc k) - x) else 0)\<^sup>2) \ d\<^sup>2" if "x \ space (lebesgue_on {a..b})" for k x using that apply (simp add: abs_mult divide_simps flip: abs_le_square_iff) apply (auto simp: abs_if zero_less_mult_iff mult_left_le) done qed have integ: "integrable (lebesgue_on {a..b}) ?g" using integrable_dominated_convergence [OF bmg bmf int_d2 tends le_d2] by auto have int: "(\k. integral\<^sup>L (lebesgue_on {a..b}) (?f k)) \ integral\<^sup>L (lebesgue_on {a..b}) ?g" using integral_dominated_convergence [OF bmg bmf int_d2 tends le_d2] by auto then obtain N where N: "\k. k \ N \ \integral\<^sup>L (lebesgue_on {a..b}) (?f k) - integral\<^sup>L (lebesgue_on {a..b}) ?g\ < e\<^sup>2" apply (simp add: lim_sequentially dist_real_def) apply (drule_tac x="e^2" in spec) using \e > 0\ by auto obtain M where "M > 0" and M: "1 / real M < b-a" using real_arch_invD [of "b-a"] by (metis \a < b\ diff_gt_0_iff_gt inverse_eq_divide neq0_conv) have *: "\integral\<^sup>L (lebesgue_on {a..b}) (?f (max M N)) - integral\<^sup>L (lebesgue_on {a..b}) ?g\ < e\<^sup>2" using N by force let ?\ = "\x. if x \ a+1/(Suc (max M N)) then ((Suc (max M N)) * d) * ((a+1/(Suc (max M N))) - x) else 0" show ?thesis proof show "continuous_on {a..b} ?\" by (rule cont) have "1 / (1 + real (max M N)) \ 1 / (real M)" by (simp add: \0 < M\ frac_le) then have "\ (b \ a+1 / (1 + real (max M N)))" using M \a < b\ \M > 0\ max.cobounded1 [of M N] by linarith then show "?\ b = 0" by simp have null_a: "{a} \ null_sets (lebesgue_on {a..b})" using \a < b\ by (simp add: null_sets_restrict_space) have "LINT x|lebesgue_on {a..b}. ?g x = 0" by (force intro: integral_eq_zero_AE AE_I' [OF null_a]) then have "l2norm {a..b} ?\ < sqrt (e^2)" unfolding l2norm_def l2product_def power2_eq_square [symmetric] apply (intro real_sqrt_less_mono) using "*" by linarith then show "l2norm {a..b} ?\ < e" using \e > 0\ by auto qed auto qed lemma square_integrable_approximate_continuous_ends: assumes f: "f square_integrable {a..b}" and "a < b" "0 < e" obtains g where "continuous_on {a..b} g" "g b = g a" "g square_integrable {a..b}" "l2norm {a..b} (\x. f x - g x) < e" proof - obtain g where contg: "continuous_on UNIV g" and "g square_integrable {a..b}" and lg: "l2norm {a..b} (\x. f x - g x) < e/2" using f \e > 0\ square_integrable_approximate_continuous by (metis (full_types) box_real(2) half_gt_zero_iff lmeasurable_cbox) obtain h where conth: "continuous_on {a..b} h" and "h a = g b - g a" "h b = 0" and lh: "l2norm {a..b} h < e/2" using integral_tweak_ends \e > 0\ by (metis \a < b\ zero_less_divide_iff zero_less_numeral) have "h square_integrable {a..b}" using \continuous_on {a..b} h\ continuous_imp_square_integrable by blast show thesis proof show "continuous_on {a..b} (\x. g x + h x)" by (blast intro: continuous_on_subset [OF contg] conth continuous_intros) then show "(\x. g x + h x) square_integrable {a..b}" using continuous_imp_square_integrable by blast show "g b + h b = g a + h a" by (simp add: \h a = g b - g a\ \h b = 0\) have "l2norm {a..b} (\x. (f x - g x) + - h x) < e" proof (rule le_less_trans [OF l2norm_triangle [of "\x. f x - g x" "{a..b}" "\x. - (h x)"]]) show "(\x. f x - g x) square_integrable {a..b}" using \g square_integrable {a..b}\ f square_integrable_diff by blast show "(\x. - h x) square_integrable {a..b}" by (simp add: \h square_integrable {a..b}\) show "l2norm {a..b} (\x. f x - g x) + l2norm {a..b} (\x. - h x) < e" using \h square_integrable {a..b}\ l2norm_neg lg lh by auto qed then show "l2norm {a..b} (\x. f x - (g x + h x)) < e" by (simp add: field_simps) qed qed subsection\Hence the main approximation result\ lemma Weierstrass_l2_trig_polynomial: assumes f: "f square_integrable {-pi..pi}" and "0 < e" obtains n a b where "l2norm {-pi..pi} (\x. f x - (\k\n. a k * sin(real k * x) + b k * cos(real k * x))) < e" proof - let ?\ = "\n a b x. \k\n. a k * sin(real k * x) + b k * cos(real k * x)" obtain g where contg: "continuous_on {-pi..pi} g" and geq: "g(-pi) = g pi" and g: "g square_integrable {-pi..pi}" and norm_fg: "l2norm {-pi..pi} (\x. f x - g x) < e/2" using \e > 0\ by (auto intro: square_integrable_approximate_continuous_ends [OF f, of "e/2"]) then obtain n a b where g_phi_less: "\x. x \ {-pi..pi} \ \g x - (?\ n a b x)\ < e/6" using \e > 0\ Weierstrass_trig_polynomial [OF contg geq, of "e/6"] by (meson zero_less_divide_iff zero_less_numeral) show thesis proof have si: "(?\ n u v) square_integrable {-pi..pi}" for n::nat and u v proof (intro square_integrable_sum continuous_imp_square_integrable) show "continuous_on {-pi..pi} (\x. u k * sin (real k * x) + v k * cos (real k * x))" if "k \ {..n}" for k using that by (intro continuous_intros) qed auto have "l2norm {-pi..pi} (\x. f x - (?\ n a b x)) = l2norm {-pi..pi} (\x. (f x - g x) + (g x - (?\ n a b x)))" by simp also have "\ \ l2norm {-pi..pi} (\x. f x - g x) + l2norm {-pi..pi} (\x. g x - (?\ n a b x))" proof (rule l2norm_triangle) show "(\x. f x - g x) square_integrable {-pi..pi}" using f g square_integrable_diff by blast show "(\x. g x - (?\ n a b x)) square_integrable {-pi..pi}" using g si square_integrable_diff by blast qed also have "\ < e" proof - have g_phi: "(\x. g x - (?\ n a b x)) square_integrable {-pi..pi}" using g si square_integrable_diff by blast have "l2norm {-pi..pi} (\x. g x - (?\ n a b x)) \ e/2" unfolding l2norm_def l2product_def power2_eq_square [symmetric] proof (rule real_le_lsqrt) show "0 \ LINT x|lebesgue_on {-pi..pi}. (g x - (?\ n a b x))\<^sup>2" by (rule integral_nonneg_AE) auto have "LINT x|lebesgue_on {-pi..pi}. (g x - (?\ n a b x))\<^sup>2 \ LINT x|lebesgue_on {-pi..pi}. (e / 6) ^ 2" proof (rule integral_mono) show "integrable (lebesgue_on {-pi..pi}) (\x. (g x - (?\ n a b x))\<^sup>2)" using g_phi square_integrable_def by auto show "integrable (lebesgue_on {-pi..pi}) (\x. (e / 6)\<^sup>2)" by (intro continuous_intros continuous_imp_integrable_real) show "(g x - (?\ n a b x))\<^sup>2 \ (e / 6)\<^sup>2" if "x \ space (lebesgue_on {-pi..pi})" for x using \e > 0\ g_phi_less that apply (simp add: abs_le_square_iff [symmetric]) using less_eq_real_def by blast qed also have "\ \ (e / 2)\<^sup>2" - using \e > 0\ pi_less_4 by (auto simp: power2_eq_square) + using \e > 0\ pi_less_4 by (auto simp: power2_eq_square measure_restrict_space) finally show "LINT x|lebesgue_on {-pi..pi}. (g x - (?\ n a b x))\<^sup>2 \ (e / 2)\<^sup>2" . qed (use \e > 0\ in auto) with norm_fg show ?thesis by auto qed finally show "l2norm {-pi..pi} (\x. f x - (?\ n a b x)) < e" . qed qed proposition Weierstrass_l2_trigonometric_set: assumes f: "f square_integrable {-pi..pi}" and "0 < e" obtains n a where "l2norm {-pi..pi} (\x. f x - (\k\n. a k * trigonometric_set k x)) < e" proof - obtain n a b where lee: "l2norm {-pi..pi} (\x. f x - (\k\n. a k * sin(real k * x) + b k * cos(real k * x))) < e" using Weierstrass_l2_trig_polynomial [OF assms] . let ?a = "\k. if k = 0 then sqrt(2 * pi) * b 0 else if even k then sqrt pi * b(k div 2) else if k \ 2 * n then sqrt pi * a((Suc k) div 2) else 0" show thesis proof have [simp]: "Suc (i * 2) \ n * 2 \ i {..k\n. b k * cos (real k * x)) = (\i\n. if i = 0 then b 0 else b i * cos (real i * x))" for x by (rule sum.cong) auto moreover have "(\k\n. a k * sin (real k * x)) = (\i\n. (if Suc (2 * i) \ 2 * n then sqrt pi * a (Suc i) * sin ((1 + real i) * x) else 0) / sqrt pi)" (is "?lhs = ?rhs") for x proof (cases "n=0") case False then obtain n' where n': "n = Suc n'" using not0_implies_Suc by blast have "?lhs = (\k = Suc 0..n. a k * sin (real k * x))" by (simp add: atMost_atLeast0 sum_shift_lb_Suc0_0) also have "\ = (\ij\n'. a (Suc j) * sin (real (Suc j) * x)) = (\i = ?rhs" by (simp add: field_simps if_distrib [of "\x. x/_"] sum.inter_restrict [where B = "{..k\n. a k * sin(real k * x) + b k * cos(real k * x)) = (\k \ Suc(2*n). ?a k * trigonometric_set k x)" for x unfolding sum.in_pairs_0 trigonometric_set_def by (simp add: sum.distrib if_distrib [of "\x. x*_"] cong: if_cong) with lee show "l2norm {-pi..pi} (\x. f x - (\k \ Suc(2*n). ?a k * trigonometric_set k x)) < e" by auto qed qed subsection\Convergence wrt the L2 norm of trigonometric Fourier series\ definition Fourier_coefficient where "Fourier_coefficient \ orthonormal_coeff {-pi..pi} trigonometric_set" lemma Fourier_series_l2: assumes "f square_integrable {-pi..pi}" shows "(\n. l2norm {-pi..pi} (\x. f x - (\i\n. Fourier_coefficient f i * trigonometric_set i x))) \ 0" proof (clarsimp simp add: lim_sequentially dist_real_def Fourier_coefficient_def) let ?h = "\n x. (\i\n. orthonormal_coeff {-pi..pi} trigonometric_set f i * trigonometric_set i x)" show "\N. \n\N. \l2norm {-pi..pi} (\x. f x - ?h n x)\ < e" if "0 < e" for e proof - obtain N a where lte: "l2norm {-pi..pi} (\x. f x - (\k\N. a k * trigonometric_set k x)) < e" using Weierstrass_l2_trigonometric_set by (meson \0 < e\ assms) show ?thesis proof (intro exI allI impI) show "\l2norm {-pi..pi} (\x. f x - ?h m x)\ < e" if "N \ m" for m proof - have "\l2norm {-pi..pi} (\x. f x - ?h m x)\ = l2norm {-pi..pi} (\x. f x - ?h m x)" proof (rule abs_of_nonneg) show "0 \ l2norm {-pi..pi} (\x. f x - ?h m x)" apply (intro l2norm_pos_le square_integrable_diff square_integrable_sum square_integrable_lmult square_integrable_trigonometric_set assms, auto) done qed also have "\ \ l2norm {-pi..pi} (\x. f x - (\k\N. a k * trigonometric_set k x))" proof - have "(\i\m. (if i \ N then a i else 0) * trigonometric_set i x) = (\i\N. a i * trigonometric_set i x)" for x using sum.inter_restrict [where A = "{..m}" and B = "{..N}", symmetric] that by (force simp: if_distrib [of "\x. x * _"] min_absorb2 cong: if_cong) moreover have "l2norm {-pi..pi} (\x. f x - ?h m x) \ l2norm {-pi..pi} (\x. f x - (\i\m. (if i \ N then a i else 0) * trigonometric_set i x))" using orthonormal_optimal_partial_sum [OF orthonormal_system_trigonometric_set square_integrable_trigonometric_set assms] by simp ultimately show ?thesis by simp qed also have "\ < e" by (rule lte) finally show ?thesis . qed qed qed qed subsection\Fourier coefficients go to 0 (weak form of Riemann-Lebesgue)\ lemma trigonometric_set_mul_absolutely_integrable: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\x. trigonometric_set n x * f x) absolutely_integrable_on {-pi..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) show "trigonometric_set n \ borel_measurable (lebesgue_on {-pi..pi})" using square_integrable_def square_integrable_trigonometric_set by blast show "bounded (trigonometric_set n ` {-pi..pi})" unfolding bounded_iff using pi_gt3 sqrt_pi_ge1 by (rule_tac x=1 in exI) (auto simp: trigonometric_set_def dist_real_def intro: order_trans [OF abs_sin_le_one] order_trans [OF abs_cos_le_one]) qed (auto simp: assms) lemma trigonometric_set_mul_integrable: "f absolutely_integrable_on {-pi..pi} \ integrable (lebesgue_on {-pi..pi}) (\x. trigonometric_set n x * f x)" using trigonometric_set_mul_absolutely_integrable by (simp add: integrable_restrict_space set_integrable_def) lemma trigonometric_set_integrable [simp]: "integrable (lebesgue_on {-pi..pi}) (trigonometric_set n)" using trigonometric_set_mul_integrable [where f = id] by simp (metis absolutely_integrable_imp_integrable fmeasurableD interval_cbox lmeasurable_cbox square_integrable_imp_absolutely_integrable square_integrable_trigonometric_set) lemma absolutely_integrable_sin_product: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\x. sin(k * x) * f x) absolutely_integrable_on {-pi..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) show "(\x. sin (k * x)) \ borel_measurable (lebesgue_on {-pi..pi})" by (metis borel_measurable_integrable integrable_sin_cx mult_commute_abs) show "bounded ((\x. sin (k * x)) ` {-pi..pi})" by (metis (mono_tags, lifting) abs_sin_le_one bounded_iff imageE real_norm_def) qed (auto simp: assms) lemma absolutely_integrable_cos_product: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\x. cos(k * x) * f x) absolutely_integrable_on {-pi..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) show "(\x. cos (k * x)) \ borel_measurable (lebesgue_on {-pi..pi})" by (metis borel_measurable_integrable integrable_cos_cx mult_commute_abs) show "bounded ((\x. cos (k * x)) ` {-pi..pi})" by (metis (mono_tags, lifting) abs_cos_le_one bounded_iff imageE real_norm_def) qed (auto simp: assms) lemma assumes "f absolutely_integrable_on {-pi..pi}" shows Fourier_products_integrable_cos: "integrable (lebesgue_on {-pi..pi}) (\x. cos(k * x) * f x)" and Fourier_products_integrable_sin: "integrable (lebesgue_on {-pi..pi}) (\x. sin(k * x) * f x)" using absolutely_integrable_cos_product absolutely_integrable_sin_product assms by (auto simp: integrable_restrict_space set_integrable_def) lemma Riemann_lebesgue_square_integrable: assumes "orthonormal_system S w" "\i. w i square_integrable S" "f square_integrable S" shows "orthonormal_coeff S w f \ 0" using Fourier_series_square_summable [OF assms, of UNIV] summable_LIMSEQ_zero by force proposition Riemann_lebesgue: assumes "f absolutely_integrable_on {-pi..pi}" shows "Fourier_coefficient f \ 0" unfolding lim_sequentially proof (intro allI impI) fix e::real assume "e > 0" then obtain g where "continuous_on UNIV g" and gabs: "g absolutely_integrable_on {-pi..pi}" and fg_e2: "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. \f x - g x\) < e/2" using absolutely_integrable_approximate_continuous [OF assms, of "e/2"] by (metis (full_types) box_real(2) half_gt_zero_iff lmeasurable_cbox) have "g square_integrable {-pi..pi}" using \continuous_on UNIV g\ continuous_imp_square_integrable continuous_on_subset by blast then have "orthonormal_coeff {-pi..pi} trigonometric_set g \ 0" using Riemann_lebesgue_square_integrable orthonormal_system_trigonometric_set square_integrable_trigonometric_set by blast with \e > 0\ obtain N where N: "\n. n \ N \ \orthonormal_coeff {-pi..pi} trigonometric_set g n\ < e/2" unfolding lim_sequentially by (metis half_gt_zero_iff norm_conv_dist real_norm_def) have "\Fourier_coefficient f n\ < e" if "n \ N" for n proof - have "\LINT x|lebesgue_on {-pi..pi}. trigonometric_set n x * g x\ < e/2" using N [OF \n \ N\] by (auto simp: orthonormal_coeff_def l2product_def) have "integrable (lebesgue_on {-pi..pi}) (\x. trigonometric_set n x * g x)" using gabs trigonometric_set_mul_integrable by blast moreover have "integrable (lebesgue_on {-pi..pi}) (\x. trigonometric_set n x * f x)" using assms trigonometric_set_mul_integrable by blast ultimately have "\(LINT x|lebesgue_on {-pi..pi}. trigonometric_set n x * g x) - (LINT x|lebesgue_on {-pi..pi}. trigonometric_set n x * f x)\ = \(LINT x|lebesgue_on {-pi..pi}. trigonometric_set n x * (g x - f x))\" by (simp add: algebra_simps flip: Bochner_Integration.integral_diff) also have "\ \ LINT x|lebesgue_on {-pi..pi}. \f x - g x\" proof (rule integral_abs_bound_integral) show "integrable (lebesgue_on {-pi..pi}) (\x. trigonometric_set n x * (g x - f x))" by (simp add: gabs assms trigonometric_set_mul_integrable) have "(\x. f x - g x) absolutely_integrable_on {-pi..pi}" using gabs assms by blast then show "integrable (lebesgue_on {-pi..pi}) (\x. \f x - g x\)" by (simp add: absolutely_integrable_imp_integrable) fix x assume "x \ space (lebesgue_on {-pi..pi})" then have "-pi \ x" "x \ pi" by auto have "\trigonometric_set n x\ \ 1" using pi_ge_two apply (simp add: trigonometric_set_def) using sqrt_pi_ge1 abs_sin_le_one order_trans abs_cos_le_one by metis then show "\trigonometric_set n x * (g x - f x)\ \ \f x - g x\" using abs_ge_zero mult_right_mono by (fastforce simp add: abs_mult abs_minus_commute) qed finally have "\(LINT x|lebesgue_on {-pi..pi}. trigonometric_set n x * g x) - (LINT x|lebesgue_on {-pi..pi}. trigonometric_set n x * f x)\ \ LINT x|lebesgue_on {-pi..pi}. \f x - g x\" . then show ?thesis using N [OF \n \ N\] fg_e2 unfolding Fourier_coefficient_def orthonormal_coeff_def l2product_def by linarith qed then show "\no. \n\no. dist (Fourier_coefficient f n) 0 < e" by auto qed lemma Riemann_lebesgue_sin: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\n. integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin(real n * x) * f x)) \ 0" unfolding lim_sequentially proof (intro allI impI) fix e::real assume "e > 0" then obtain N where N: "\n. n \ N \ \Fourier_coefficient f n\ < e/4" using Riemann_lebesgue [OF assms] unfolding lim_sequentially by (metis norm_conv_dist real_norm_def zero_less_divide_iff zero_less_numeral) have "\LINT x|lebesgue_on {-pi..pi}. sin (real n * x) * f x\ < e" if "n > N" for n using that proof (induction n) case (Suc n) have "\Fourier_coefficient f(Suc (2 * n))\ < e/4" using N Suc.prems by auto then have "\LINT x|lebesgue_on {-pi..pi}. sin ((1 + real n) * x) * f x\ < sqrt pi * e / 4" by (simp add: Fourier_coefficient_def orthonormal_coeff_def trigonometric_set_def l2product_def field_simps) also have "\ \ e" using \0 < e\ pi_less_4 real_sqrt_less_mono by (fastforce simp add: field_simps) finally show ?case by simp qed auto then show "\no. \n\no. dist (LINT x|lebesgue_on {-pi..pi}. sin (real n * x) * f x) 0 < e" by (metis (full_types) le_neq_implies_less less_add_same_cancel2 less_trans norm_conv_dist real_norm_def zero_less_one) qed lemma Riemann_lebesgue_cos: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\n. integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(real n * x) * f x)) \ 0" unfolding lim_sequentially proof (intro allI impI) fix e::real assume "e > 0" then obtain N where N: "\n. n \ N \ \Fourier_coefficient f n\ < e/4" using Riemann_lebesgue [OF assms] unfolding lim_sequentially by (metis norm_conv_dist real_norm_def zero_less_divide_iff zero_less_numeral) have "\LINT x|lebesgue_on {-pi..pi}. cos (real n * x) * f x\ < e" if "n > N" for n using that proof (induction n) case (Suc n) have eq: "(x * 2 + x * (real n * 2)) / 2 = x + x * (real n)" for x by simp have "\Fourier_coefficient f(2*n + 2)\ < e/4" using N Suc.prems by auto then have "\LINT x|lebesgue_on {-pi..pi}. f x * cos (x + x * (real n))\ < sqrt pi * e / 4" by (simp add: Fourier_coefficient_def orthonormal_coeff_def trigonometric_set_def l2product_def field_simps eq) also have "\ \ e" using \0 < e\ pi_less_4 real_sqrt_less_mono by (fastforce simp add: field_simps) finally show ?case by (simp add: field_simps) qed auto then show "\no. \n\no. dist (LINT x|lebesgue_on {-pi..pi}. cos (real n * x) * f x) 0 < e" by (metis (full_types) le_neq_implies_less less_add_same_cancel2 less_trans norm_conv_dist real_norm_def zero_less_one) qed lemma Riemann_lebesgue_sin_half: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\n. LINT x|lebesgue_on {-pi..pi}. sin ((real n + 1/2) * x) * f x) \ 0" proof (simp add: algebra_simps sin_add) let ?INT = "integral\<^sup>L (lebesgue_on {-pi..pi})" let ?f = "(\n. ?INT (\x. sin(n * x) * cos(1/2 * x) * f x) + ?INT (\x. cos(n * x) * sin(1/2 * x) * f x))" show "(\n. ?INT (\x. f x * (cos (x * real n) * sin (x/2)) + f x * (sin (x * real n) * cos (x/2)))) \ 0" proof (rule Lim_transform_eventually) have sin: "(\x. sin (1/2 * x) * f x) absolutely_integrable_on {-pi..pi}" by (intro absolutely_integrable_sin_product assms) have cos: "(\x. cos (1/2 * x) * f x) absolutely_integrable_on {-pi..pi}" by (intro absolutely_integrable_cos_product assms) show "\\<^sub>F n in sequentially. ?f n = ?INT (\x. f x * (cos (x * real n) * sin (x/2)) + f x * (sin (x * real n) * cos (x/2)))" unfolding mult.assoc apply (rule eventuallyI) apply (subst Bochner_Integration.integral_add [symmetric]) apply (safe intro!: cos absolutely_integrable_sin_product sin absolutely_integrable_cos_product absolutely_integrable_imp_integrable) apply (auto simp: algebra_simps) done have "?f \ 0 + 0" unfolding mult.assoc by (intro tendsto_add Riemann_lebesgue_sin Riemann_lebesgue_cos sin cos) then show "?f \ (0::real)" by simp qed qed lemma Fourier_sum_limit_pair: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\n. \k\2 * n. Fourier_coefficient f k * trigonometric_set k t) \ l \ (\n. \k\n. Fourier_coefficient f k * trigonometric_set k t) \ l" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs unfolding lim_sequentially dist_real_def proof (intro allI impI) fix e::real assume "e > 0" then obtain N1 where N1: "\n. n \ N1 \ \Fourier_coefficient f n\ < e/2" using Riemann_lebesgue [OF assms] unfolding lim_sequentially by (metis norm_conv_dist real_norm_def zero_less_divide_iff zero_less_numeral) obtain N2 where N2: "\n. n \ N2 \ \(\k\2 * n. Fourier_coefficient f k * trigonometric_set k t) - l\ < e/2" using L unfolding lim_sequentially dist_real_def by (meson \0 < e\ half_gt_zero_iff) show "\no. \n\no. \(\k\n. Fourier_coefficient f k * trigonometric_set k t) - l\ < e" proof (intro exI allI impI) fix n assume n: "N1 + 2*N2 + 1 \ n" then have "n \ N1" "n \ N2" "n div 2 \ N2" by linarith+ consider "n = 2 * (n div 2)" | "n = Suc(2 * (n div 2))" by linarith then show "\(\k\n. Fourier_coefficient f k * trigonometric_set k t) - l\ < e" proof cases case 1 show ?thesis apply (subst 1) using N2 [OF \n div 2 \ N2\] by linarith next case 2 have "\(\k\2 * (n div 2). Fourier_coefficient f k * trigonometric_set k t) - l\ < e/2" using N2 [OF \n div 2 \ N2\] by linarith moreover have "\Fourier_coefficient f(Suc (2 * (n div 2))) * trigonometric_set (Suc (2 * (n div 2))) t\ < (e/2) * 1" proof - have "\trigonometric_set (Suc (2 * (n div 2))) t\ \ 1" apply (simp add: trigonometric_set) using abs_sin_le_one sqrt_pi_ge1 by (blast intro: order_trans) moreover have "\Fourier_coefficient f(Suc (2 * (n div 2)))\ < e/2" using N1 \N1 \ n\ by auto ultimately show ?thesis unfolding abs_mult by (meson abs_ge_zero le_less_trans mult_left_mono real_mult_less_iff1 zero_less_one) qed ultimately show ?thesis apply (subst 2) unfolding sum.atMost_Suc by linarith qed qed qed next assume ?rhs then show ?lhs unfolding lim_sequentially by (auto simp: elim!: imp_forward ex_forward) qed subsection\Express Fourier sum in terms of the special expansion at the origin\ lemma Fourier_sum_0: "(\k \ n. Fourier_coefficient f k * trigonometric_set k 0) = (\k \ n div 2. Fourier_coefficient f(2*k) * trigonometric_set (2*k) 0)" (is "?lhs = ?rhs") proof - have "?lhs = (\k = 2 * 0.. Suc (2 * (n div 2)). Fourier_coefficient f k * trigonometric_set k 0)" proof (rule sum.mono_neutral_left) show "\i\{2 * 0..Suc (2 * (n div 2))} - {..n}. Fourier_coefficient f i * trigonometric_set i 0 = 0" proof clarsimp fix i assume "i \ Suc (2 * (n div 2))" "\ i \ n" then have "i = Suc n" "even n" by presburger+ moreover assume "trigonometric_set i 0 \ 0" ultimately show "Fourier_coefficient f i = 0" by (simp add: trigonometric_set_def) qed qed auto also have "\ = ?rhs" unfolding sum.in_pairs by (simp add: trigonometric_set atMost_atLeast0) finally show ?thesis . qed lemma Fourier_sum_0_explicit: "(\k\n. Fourier_coefficient f k * trigonometric_set k 0) = (Fourier_coefficient f 0 / sqrt 2 + (\k = 1..n div 2. Fourier_coefficient f(2*k))) / sqrt pi" (is "?lhs = ?rhs") proof - have "(\k=0..n. Fourier_coefficient f k * trigonometric_set k 0) = Fourier_coefficient f 0 * trigonometric_set 0 0 + (\k = 1..n. Fourier_coefficient f k * trigonometric_set k 0)" by (simp add: Fourier_sum_0 sum.atLeast_Suc_atMost) also have "\ = ?rhs" proof - have "Fourier_coefficient f 0 * trigonometric_set 0 0 = Fourier_coefficient f 0 / (sqrt 2 * sqrt pi)" by (simp add: real_sqrt_mult trigonometric_set(1)) moreover have "(\k = Suc 0..n. Fourier_coefficient f k * trigonometric_set k 0) = (\k = Suc 0..n div 2. Fourier_coefficient f(2*k)) / sqrt pi" proof (induction n) case (Suc n) show ?case by (simp add: Suc) (auto simp: trigonometric_set_def field_simps) qed auto ultimately show ?thesis by (simp add: add_divide_distrib) qed finally show ?thesis by (simp add: atMost_atLeast0) qed lemma Fourier_sum_0_integrals: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\k\n. Fourier_coefficient f k * trigonometric_set k 0) = (integral\<^sup>L (lebesgue_on {-pi..pi}) f / 2 + (\k = Suc 0..n div 2. integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(k * x) * f x))) / pi" proof - have "integral\<^sup>L (lebesgue_on {-pi..pi}) f / (sqrt (2 * pi) * sqrt 2 * sqrt pi) = integral\<^sup>L (lebesgue_on {-pi..pi}) f / (2 * pi)" by (simp add: algebra_simps real_sqrt_mult) moreover have "(\k = Suc 0..n div 2. LINT x|lebesgue_on {-pi..pi}. trigonometric_set (2*k) x * f x) / sqrt pi = (\k = Suc 0..n div 2. LINT x|lebesgue_on {-pi..pi}. cos (k * x) * f x) / pi" by (simp add: trigonometric_set_def sum_divide_distrib) ultimately show ?thesis unfolding Fourier_sum_0_explicit by (simp add: Fourier_coefficient_def orthonormal_coeff_def l2product_def trigonometric_set add_divide_distrib) qed lemma Fourier_sum_0_integral: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\k\n. Fourier_coefficient f k * trigonometric_set k 0) = integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. (1/2 + (\k = Suc 0..n div 2. cos(k * x))) * f x) / pi" proof - note * = Fourier_products_integrable_cos [OF assms] have "integrable (lebesgue_on {-pi..pi}) (\x. \n = Suc 0..n div 2. f x * cos (x * real n))" using * by (force simp: mult_ac) moreover have "integral\<^sup>L (lebesgue_on {-pi..pi}) f / 2 + (\k = Suc 0..n div 2. LINT x|lebesgue_on {-pi..pi}. f x * cos (x * real k)) = (LINT x|lebesgue_on {-pi..pi}. f x / 2) + (LINT x|lebesgue_on {-pi..pi}. (\n = Suc 0..n div 2. f x * cos (x * real n)))" proof (subst Bochner_Integration.integral_sum) show "integrable (lebesgue_on {-pi..pi}) (\x. f x * cos (x * real i))" if "i \ {Suc 0..n div 2}" for i using that * [of i] by (auto simp: Bochner_Integration.integral_sum mult_ac) qed auto ultimately show ?thesis using Fourier_products_integrable_cos [OF assms] absolutely_integrable_imp_integrable [OF assms] by (simp add: Fourier_sum_0_integrals sum_distrib_left assms algebra_simps) qed subsection\How Fourier coefficients behave under addition etc\ lemma Fourier_coefficient_add: assumes "f absolutely_integrable_on {-pi..pi}" "g absolutely_integrable_on {-pi..pi}" shows "Fourier_coefficient (\x. f x + g x) i = Fourier_coefficient f i + Fourier_coefficient g i" using assms trigonometric_set_mul_integrable by (simp add: Fourier_coefficient_def orthonormal_coeff_def l2product_def distrib_left) lemma Fourier_coefficient_minus: assumes "f absolutely_integrable_on {-pi..pi}" shows "Fourier_coefficient (\x. - f x) i = - Fourier_coefficient f i" using assms trigonometric_set_mul_integrable by (simp add: Fourier_coefficient_def orthonormal_coeff_def l2product_def) lemma Fourier_coefficient_diff: assumes f: "f absolutely_integrable_on {-pi..pi}" and g: "g absolutely_integrable_on {-pi..pi}" shows "Fourier_coefficient (\x. f x - g x) i = Fourier_coefficient f i - Fourier_coefficient g i" proof - have mg: "(\x. - g x) absolutely_integrable_on {-pi..pi}" using g by (simp add: absolutely_integrable_measurable_real) show ?thesis using Fourier_coefficient_add [OF f mg] Fourier_coefficient_minus [OF g] by simp qed lemma Fourier_coefficient_const: "Fourier_coefficient (\x. c) i = (if i = 0 then c * sqrt(2 * pi) else 0)" - by (auto simp: Fourier_coefficient_def orthonormal_coeff_def l2product_def trigonometric_set_def divide_simps) + by (auto simp: Fourier_coefficient_def orthonormal_coeff_def l2product_def trigonometric_set_def divide_simps measure_restrict_space) lemma Fourier_offset_term: fixes f :: "real \ real" assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "Fourier_coefficient (\x. f(x+t)) (2 * n + 2) * trigonometric_set (2 * n + 2) 0 = Fourier_coefficient f(2 * n+1) * trigonometric_set (2 * n+1) t + Fourier_coefficient f(2 * n + 2) * trigonometric_set (2 * n + 2) t" proof - have eq: "(2 + 2 * real n) * x / 2 = (1 + real n) * x" for x by (simp add: divide_simps) have 1: "integrable (lebesgue_on {-pi..pi}) (\x. f x * (cos (x + x * n) * cos (t + t * n)))" using Fourier_products_integrable_cos [of f "Suc n"] f by (simp add: algebra_simps) (simp add: mult.assoc [symmetric]) have 2: "integrable (lebesgue_on {-pi..pi}) (\x. f x * (sin (x + x * n) * sin (t + t * n)))" using Fourier_products_integrable_sin [of f "Suc n"] f by (simp add: algebra_simps) (simp add: mult.assoc [symmetric]) have "has_bochner_integral (lebesgue_on {-pi..pi}) (\x. cos (real (Suc n) * (x + t - t)) * f(x + t)) (LINT x|lebesgue_on {-pi..pi}. cos (real (Suc n) * (x - t)) * f x)" proof (rule has_integral_periodic_offset) have "(\x. cos (real (Suc n) * (x - t)) * f x) absolutely_integrable_on {-pi..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) show "(\x. cos (real (Suc n) * (x - t))) \ borel_measurable (lebesgue_on {-pi..pi})" by (intro continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto show "bounded ((\x. cos (real (Suc n) * (x - t))) ` {-pi..pi})" by (rule boundedI [where B=1]) auto qed (use f in auto) then show "has_bochner_integral (lebesgue_on {-pi..pi}) (\x. cos ((Suc n) * (x - t)) * f x) (LINT x|lebesgue_on {-pi..pi}. cos (real (Suc n) * (x - t)) * f x)" by (simp add: has_bochner_integral_integrable integrable_restrict_space set_integrable_def) next fix x show "cos (real (Suc n) * (x + (pi - - pi) - t)) * f(x + (pi - - pi)) = cos (real (Suc n) * (x - t)) * f x" using periodic cos.plus_of_nat [of "(Suc n) * (x - t)" "Suc n"] by (simp add: algebra_simps) qed then have "has_bochner_integral (lebesgue_on {-pi..pi}) (\x. cos (real (Suc n) * x) * f(x + t)) (LINT x|lebesgue_on {-pi..pi}. cos (real (Suc n) * (x - t)) * f x)" by simp then have "LINT x|lebesgue_on {-pi..pi}. cos ((Suc n) * x) * f(x+t) = LINT x|lebesgue_on {-pi..pi}. cos ((Suc n) * (x - t)) * f x" using has_bochner_integral_integral_eq by blast also have "\ = LINT x|lebesgue_on {-pi..pi}. cos ((Suc n) * x - ((Suc n) * t)) * f x" by (simp add: algebra_simps) also have "\ = cos ((Suc n) * t) * (LINT x|lebesgue_on {-pi..pi}. cos ((Suc n) * x) * f x) + sin ((Suc n) * t) * (LINT x|lebesgue_on {-pi..pi}. sin ((Suc n) * x) * f x)" by (simp add: cos_diff algebra_simps 1 2 flip: integral_mult_left_zero integral_mult_right_zero) finally have "LINT x|lebesgue_on {-pi..pi}. cos ((Suc n) * x) * f(x+t) = cos ((Suc n) * t) * (LINT x|lebesgue_on {-pi..pi}. cos ((Suc n) * x) * f x) + sin ((Suc n) * t) * (LINT x|lebesgue_on {-pi..pi}. sin ((Suc n) * x) * f x)" . then show ?thesis unfolding Fourier_coefficient_def orthonormal_coeff_def trigonometric_set_def by (simp add: eq) (simp add: divide_simps algebra_simps l2product_def) qed lemma Fourier_sum_offset: fixes f :: "real \ real" assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\k\2*n. Fourier_coefficient f k * trigonometric_set k t) = (\k\2*n. Fourier_coefficient (\x. f(x+t)) k * trigonometric_set k 0)" (is "?lhs = ?rhs") proof - have "?lhs = Fourier_coefficient f 0 * trigonometric_set 0 t + (\k = Suc 0..2*n. Fourier_coefficient f k * trigonometric_set k t)" by (simp add: atMost_atLeast0 sum.atLeast_Suc_atMost) moreover have "(\k = Suc 0..2*n. Fourier_coefficient f k * trigonometric_set k t) = (\k = Suc 0..2*n. Fourier_coefficient (\x. f(x+t)) k * trigonometric_set k 0)" proof (cases n) case (Suc n') have "(\k = Suc 0..2 * n. Fourier_coefficient f k * trigonometric_set k t) = (\k = Suc 0.. Suc (Suc (2 * n')). Fourier_coefficient f k * trigonometric_set k t)" by (simp add: Suc) also have "\ = (\k \ Suc (2 * n'). Fourier_coefficient f(Suc k) * trigonometric_set (Suc k) t)" unfolding atMost_atLeast0 using sum.shift_bounds_cl_Suc_ivl by blast also have "\ = (\k \ Suc (2 * n'). Fourier_coefficient (\x. f(x+t)) (Suc k) * trigonometric_set (Suc k) 0)" unfolding sum.in_pairs_0 proof (rule sum.cong [OF refl]) show "Fourier_coefficient f(Suc (2 * k)) * trigonometric_set (Suc (2 * k)) t + Fourier_coefficient f(Suc (Suc (2 * k))) * trigonometric_set (Suc (Suc (2 * k))) t = Fourier_coefficient (\x. f(x + t)) (Suc (2 * k)) * trigonometric_set (Suc (2 * k)) 0 + Fourier_coefficient (\x. f(x + t)) (Suc (Suc (2 * k))) * trigonometric_set (Suc (Suc (2 * k))) 0" if "k \ {..n'}" for k using that Fourier_offset_term [OF assms, of t ] by (auto simp: trigonometric_set_def) qed also have "\ = (\k = Suc 0..Suc (Suc (2 * n')). Fourier_coefficient (\x. f(x+t)) k * trigonometric_set k 0)" by (simp only: atMost_atLeast0 sum.shift_bounds_cl_Suc_ivl) also have "\ = (\k = Suc 0..2*n. Fourier_coefficient (\x. f(x+t)) k * trigonometric_set k 0)" by (simp add: Suc) finally show ?thesis . qed auto moreover have "?rhs = Fourier_coefficient (\x. f(x+t)) 0 * trigonometric_set 0 0 + (\k = Suc 0..2*n. Fourier_coefficient (\x. f(x+t)) k * trigonometric_set k 0)" by (simp add: atMost_atLeast0 sum.atLeast_Suc_atMost) moreover have "Fourier_coefficient f 0 * trigonometric_set 0 t = Fourier_coefficient (\x. f(x+t)) 0 * trigonometric_set 0 0" by (simp add: Fourier_coefficient_def orthonormal_coeff_def trigonometric_set_def l2product_def integral_periodic_offset periodic) ultimately show ?thesis by simp qed lemma Fourier_sum_offset_unpaired: fixes f :: "real \ real" assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\k\2*n. Fourier_coefficient f k * trigonometric_set k t) = (\k\n. Fourier_coefficient (\x. f(x+t)) (2*k) * trigonometric_set (2*k) 0)" (is "?lhs = ?rhs") proof - have "?lhs = (\k\n. Fourier_coefficient (\x. f(x+t)) (2*k) * trigonometric_set (2*k) 0 + Fourier_coefficient (\x. f(x+t)) (Suc (2*k)) * trigonometric_set (Suc (2*k)) 0)" apply (simp add: assms Fourier_sum_offset) apply (subst sum.in_pairs_0 [symmetric]) by (simp add: trigonometric_set) also have "\ = ?rhs" by (auto simp: trigonometric_set) finally show ?thesis . qed subsection\Express partial sums using Dirichlet kernel\ definition Dirichlet_kernel where "Dirichlet_kernel \ \n x. if x = 0 then real n + 1/2 else sin((real n + 1/2) * x) / (2 * sin(x/2))" lemma Dirichlet_kernel_0 [simp]: "\x\ < 2 * pi \ Dirichlet_kernel 0 x = 1/2" by (auto simp: Dirichlet_kernel_def sin_zero_iff) lemma Dirichlet_kernel_minus [simp]: "Dirichlet_kernel n (-x) = Dirichlet_kernel n x" by (auto simp: Dirichlet_kernel_def) lemma Dirichlet_kernel_continuous_strong: "continuous_on {-(2 * pi)<..<2 * pi} (Dirichlet_kernel n)" proof - have "isCont (Dirichlet_kernel n) z" if "-(2 * pi) < z" "z < 2 * pi" for z proof (cases "z=0") case True have *: "(\x. sin ((n + 1/2) * x) / (2 * sin (x/2))) \0 \ real n + 1/2" by real_asymp show ?thesis unfolding Dirichlet_kernel_def continuous_at True apply (rule Lim_transform_eventually [where f = "\x. sin((n + 1/2) * x) / (2 * sin(x/2))"]) apply (auto simp: * eventually_at_filter) done next case False have "continuous (at z) (\x. sin((real n + 1/2) * x) / (2 * sin(x/2)))" using that False by (intro continuous_intros) (auto simp: sin_zero_iff) moreover have "\\<^sub>F x in nhds z. x \ 0" using False t1_space_nhds by blast ultimately show ?thesis using False by (force simp: Dirichlet_kernel_def continuous_at eventually_at_filter elim: Lim_transform_eventually) qed then show ?thesis by (simp add: continuous_on_eq_continuous_at) qed lemma Dirichlet_kernel_continuous: "continuous_on {-pi..pi} (Dirichlet_kernel n)" apply (rule continuous_on_subset [OF Dirichlet_kernel_continuous_strong], clarsimp) using pi_gt_zero by linarith lemma absolutely_integrable_mult_Dirichlet_kernel: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\x. Dirichlet_kernel n x * f x) absolutely_integrable_on {-pi..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) show "Dirichlet_kernel n \ borel_measurable (lebesgue_on {-pi..pi})" by (simp add: Dirichlet_kernel_continuous continuous_imp_measurable_on_sets_lebesgue) have "compact (Dirichlet_kernel n ` {-pi..pi})" by (auto simp: compact_continuous_image [OF Dirichlet_kernel_continuous]) then show "bounded (Dirichlet_kernel n ` {-pi..pi})" using compact_imp_bounded by blast qed (use assms in auto) lemma cosine_sum_lemma: "(1/2 + (\k = Suc 0..n. cos(real k * x))) * sin(x/2) = sin((real n + 1/2) * x) / 2" proof - consider "n=0" | "n \ 1" by linarith then show ?thesis proof cases case 2 then have "(\k = Suc 0..n. (sin (real k * x + x/2) * 2 - sin (real k * x - x/2) * 2) / 2) = sin (real n * x + x/2) - sin (x/2)" proof (induction n) case (Suc n) show ?case proof (cases "n=0") case False with Suc show ?thesis by (simp add: divide_simps algebra_simps) qed auto qed auto then show ?thesis by (simp add: distrib_right sum_distrib_right cos_times_sin) qed auto qed lemma Dirichlet_kernel_cosine_sum: assumes "\x\ < 2 * pi" shows "Dirichlet_kernel n x = 1/2 + (\k = Suc 0..n. cos(real k * x))" proof - have "sin ((real n + 1/2) * x) / (2 * sin (x/2)) = 1/2 + (\k = Suc 0..n. cos (real k * x))" if "x \ 0" proof - have "sin(x/2) \ 0" using assms that by (auto simp: sin_zero_iff) then show ?thesis using cosine_sum_lemma [of x n] by (simp add: divide_simps) qed then show ?thesis by (auto simp: Dirichlet_kernel_def) qed lemma integrable_Dirichlet_kernel: "integrable (lebesgue_on {-pi..pi}) (Dirichlet_kernel n)" using Dirichlet_kernel_continuous continuous_imp_integrable_real by blast lemma integral_Dirichlet_kernel [simp]: "integral\<^sup>L (lebesgue_on {-pi..pi}) (Dirichlet_kernel n) = pi" proof - have "integral\<^sup>L (lebesgue_on {-pi..pi}) (Dirichlet_kernel n) = LINT x|lebesgue_on {-pi..pi}. 1/2 + (\k = Suc 0..n. cos (k * x))" using pi_ge_two by (force intro: Bochner_Integration.integral_cong [OF refl Dirichlet_kernel_cosine_sum]) also have "\ = (LINT x|lebesgue_on {-pi..pi}. 1/2) + (LINT x|lebesgue_on {-pi..pi}. (\k = Suc 0..n. cos (k * x)))" proof (rule Bochner_Integration.integral_add) show "integrable (lebesgue_on {-pi..pi}) (\x. \k = Suc 0..n. cos (real k * x))" by (rule Bochner_Integration.integrable_sum) (metis integrable_cos_cx mult_commute_abs) qed auto also have "\ = pi" proof - have "\i. \Suc 0 \ i; i \ n\ \ integrable (lebesgue_on {-pi..pi}) (\x. cos (real i * x))" by (metis integrable_cos_cx mult_commute_abs) then have "LINT x|lebesgue_on {-pi..pi}. (\k = Suc 0..n. cos (real k * x)) = 0" by (simp add: Bochner_Integration.integral_sum) then show ?thesis - by auto + by (auto simp: measure_restrict_space) qed finally show ?thesis . qed lemma integral_Dirichlet_kernel_half [simp]: "integral\<^sup>L (lebesgue_on {0..pi}) (Dirichlet_kernel n) = pi/2" proof - have "integral\<^sup>L (lebesgue_on {- pi..0}) (Dirichlet_kernel n) + integral\<^sup>L (lebesgue_on {0..pi}) (Dirichlet_kernel n) = pi" using integral_combine [OF integrable_Dirichlet_kernel] integral_Dirichlet_kernel by simp moreover have "integral\<^sup>L (lebesgue_on {- pi..0}) (Dirichlet_kernel n) = integral\<^sup>L (lebesgue_on {0..pi}) (Dirichlet_kernel n)" using integral_reflect_real [of pi 0 "Dirichlet_kernel n"] by simp ultimately show ?thesis by simp qed lemma Fourier_sum_offset_Dirichlet_kernel: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\k\2*n. Fourier_coefficient f k * trigonometric_set k t) = integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. Dirichlet_kernel n x * f(x+t)) / pi" (is "?lhs = ?rhs") proof - have ft: "(\x. f(x+t)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset [OF f, of t] periodic by simp have "?lhs = (\k=0..n. Fourier_coefficient (\x. f(x+t)) (2*k) * trigonometric_set (2*k) 0)" using Fourier_sum_offset_unpaired assms atMost_atLeast0 by auto also have "\ = Fourier_coefficient (\x. f(x+t)) 0 / sqrt (2 * pi) + (\k = Suc 0..n. Fourier_coefficient (\x. f(x+t)) (2*k) / sqrt pi)" by (simp add: sum.atLeast_Suc_atMost trigonometric_set_def) also have "\ = (LINT x|lebesgue_on {-pi..pi}. f(x+t)) / (2 * pi) + (\k = Suc 0..n. (LINT x|lebesgue_on {-pi..pi}. cos (real k * x) * f(x+t)) / pi)" by (simp add: Fourier_coefficient_def orthonormal_coeff_def trigonometric_set_def l2product_def) also have "\ = LINT x|lebesgue_on {-pi..pi}. f(x+t) / (2 * pi) + (\k = Suc 0..n. (cos (real k * x) * f(x+t)) / pi)" using Fourier_products_integrable_cos [OF ft] absolutely_integrable_imp_integrable [OF ft] by simp also have "\ = (LINT x|lebesgue_on {-pi..pi}. f(x+t) / 2 + (\k = Suc 0..n. cos (real k * x) * f(x+t))) / pi" by (simp add: divide_simps sum_distrib_right mult.assoc) also have "\ = ?rhs" proof - have "LINT x|lebesgue_on {-pi..pi}. f(x + t) / 2 + (\k = Suc 0..n. cos (real k * x) * f(x + t)) = LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * f(x + t)" proof - have eq: "f(x+t) / 2 + (\k = Suc 0..n. cos (real k * x) * f(x + t)) = Dirichlet_kernel n x * f(x + t)" if "- pi \ x" "x \ pi" for x proof (cases "x = 0") case False then have "sin (x/2) \ 0" using that by (auto simp: sin_zero_iff) then have "f(x + t) * (1/2 + (\k = Suc 0..n. cos(real k * x))) = f(x + t) * sin((real n + 1/2) * x) / 2 / sin(x/2)" using cosine_sum_lemma [of x n] by (simp add: divide_simps) then show ?thesis by (simp add: Dirichlet_kernel_def False field_simps sum_distrib_left) qed (simp add: Dirichlet_kernel_def algebra_simps) show ?thesis by (rule Bochner_Integration.integral_cong [OF refl]) (simp add: eq) qed then show ?thesis by simp qed finally show ?thesis . qed lemma Fourier_sum_limit_Dirichlet_kernel: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "((\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ l) \ (\n. LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * f(x + t)) \ pi * l" (is "?lhs = ?rhs") proof - have "?lhs \ (\n. (LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * f(x + t)) / pi) \ l" by (simp add: Fourier_sum_limit_pair [OF f, symmetric] Fourier_sum_offset_Dirichlet_kernel assms) also have "\ \ (\k. (LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel k x * f(x + t)) * inverse pi) \ pi * l * inverse pi" by (auto simp: divide_simps) also have "\ \ ?rhs" - using real_tendsto_mult_right_iff [of "inverse pi", symmetric] by auto + using tendsto_mult_right_iff [of "inverse pi", symmetric] by auto finally show ?thesis . qed subsection\A directly deduced sufficient condition for convergence at a point\ lemma simple_Fourier_convergence_periodic: assumes f: "f absolutely_integrable_on {-pi..pi}" and ft: "(\x. (f(x+t) - f t) / sin(x/2)) absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ f t" proof - let ?\ = "\n. \k\n. Fourier_coefficient (\x. f x - f t) k * trigonometric_set k t" let ?\ = "\n. LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * (f(x + t) - f t)" have fft: "(\x. f x - f t) absolutely_integrable_on {-pi..pi}" by (simp add: f absolutely_integrable_measurable_real) have fxt: "(\x. f(x + t)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset assms by auto have *: "?\ \ 0 \ ?\ \ 0" by (simp add: Fourier_sum_limit_Dirichlet_kernel fft periodic) moreover have "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t) - f t) \ 0" if "?\ \ 0" proof (rule Lim_transform_eventually [OF that eventually_sequentiallyI]) show "(\k\n. Fourier_coefficient (\x. f x - f t) k * trigonometric_set k t) = (\k\n. Fourier_coefficient f k * trigonometric_set k t) - f t" if "Suc 0 \ n" for n proof - have "(\k = Suc 0..n. Fourier_coefficient (\x. f x - f t) k * trigonometric_set k t) = (\k = Suc 0..n. Fourier_coefficient f k * trigonometric_set k t)" proof (rule sum.cong [OF refl]) fix k assume k: "k \ {Suc 0..n}" then have [simp]: "integral\<^sup>L (lebesgue_on {-pi..pi}) (trigonometric_set k) = 0" by (auto simp: trigonometric_set_def) show "Fourier_coefficient (\x. f x - f t) k * trigonometric_set k t = Fourier_coefficient f k * trigonometric_set k t" using k unfolding Fourier_coefficient_def orthonormal_coeff_def l2product_def by (simp add: right_diff_distrib f absolutely_integrable_measurable_real trigonometric_set_mul_integrable) qed moreover have "Fourier_coefficient (\x. f x - f t) 0 * trigonometric_set 0 t = Fourier_coefficient f 0 * trigonometric_set 0 t - f t" using f unfolding Fourier_coefficient_def orthonormal_coeff_def l2product_def - by (auto simp: divide_simps trigonometric_set absolutely_integrable_imp_integrable) + by (auto simp: divide_simps trigonometric_set absolutely_integrable_imp_integrable measure_restrict_space) ultimately show ?thesis by (simp add: sum.atLeast_Suc_atMost atMost_atLeast0) qed qed moreover have "?\ \ 0" proof - have eq: "\n. ?\ n = integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin((n + 1/2) * x) * ((f(x+t) - f t) / sin(x/2) / 2))" unfolding Dirichlet_kernel_def by (rule Bochner_Integration.integral_cong [OF refl]) (auto simp: divide_simps sin_zero_iff) show ?thesis unfolding eq by (intro ft set_integrable_divide Riemann_lebesgue_sin_half) qed ultimately show ?thesis by (simp add: LIM_zero_cancel) qed subsection\A more natural sufficient Hoelder condition at a point\ lemma bounded_inverse_sin_half: assumes "d > 0" obtains B where "B>0" "\x. x \ ({-pi..pi} - {-d<.. \inverse (sin (x/2))\ \ B" proof - have "bounded ((\x. inverse (sin (x/2))) ` ({-pi..pi} - {- d<..x \ {-pi..pi}; x \ 0\ \ sin (x/2) \ 0" for x using \0 < d\ by (auto simp: sin_zero_iff) then show "continuous_on ({-pi..pi} - {-d<..x. inverse (sin (x/2)))" using \0 < d\ by (intro continuous_intros) auto show "compact ({-pi..pi} - {-d<.. 0" "a > 0" and ft: "\x. \x-t\ < d \ \f x - f t\ \ M * \x-t\ powr a" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ f t" proof (rule simple_Fourier_convergence_periodic [OF f]) have half: "((\x. sin(x/2)) has_real_derivative 1/2) (at 0)" by (rule derivative_eq_intros refl | force)+ then obtain e0::real where "e0 > 0" and e0: "\x. \x \ 0; \x\ < e0\ \ \sin (x/2) / x - 1/2\ < 1/4" apply (simp add: DERIV_def Lim_at dist_real_def) apply (drule_tac x="1/4" in spec, auto) done obtain e where "e > 0" and e: "\x. \x\ < e \ \(f(x+t) - f t) / sin (x/2)\ \ 4 * (\M\ * \x\ powr (a-1))" proof show "min d e0 > 0" using \d > 0\ \e0 > 0\ by auto next fix x assume x: "\x\ < min d e0" show "\(f(x + t) - f t) / sin (x/2)\ \ 4 * (\M\ * \x\ powr (a - 1))" proof (cases "sin(x/2) = 0 \ x=0") case False have eq: "\(f(x + t) - f t) / sin (x/2)\ = \inverse (sin (x/2) / x)\ * (\f(x + t) - f t\ / \x\)" by simp show ?thesis unfolding eq proof (rule mult_mono) have "\sin (x/2) / x - 1/2\ < 1/4" using e0 [of x] x False by force then have "1/4 \ \sin (x/2) / x\" by (simp add: abs_if field_simps split: if_split_asm) then show "\inverse (sin (x/2) / x)\ \ 4" using False by (simp add: field_simps) have "\f(x + t) - f t\ / \x\ \ M * \x + t - t\ powr (a - 1)" using ft [of "x+t"] x by (auto simp: divide_simps algebra_simps Transcendental.powr_mult_base) also have "\ \ \M\ * \x\ powr (a - 1)" by (simp add: mult_mono) finally show "\f(x + t) - f t\ / \x\ \ \M\ * \x\ powr (a - 1)" . qed auto qed auto qed obtain B where "B>0" and B: "\x. x \ ({-pi..pi} - {- e<.. \inverse (sin (x/2))\ \ B" using bounded_inverse_sin_half [OF \e > 0\] by metis let ?g = "\x. max (B * \f(x + t) - f t\) ((4 * \M\) * \x\ powr (a - 1))" show "(\x. (f(x + t) - f t) / sin (x/2)) absolutely_integrable_on {-pi..pi}" proof (rule measurable_bounded_by_integrable_imp_absolutely_integrable) have fxt: "(\x. f(x + t)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset assms by auto show "(\x. (f(x + t) - f t) / sin (x/2)) \ borel_measurable (lebesgue_on {-pi..pi})" proof (intro measurable) show "(\x. f(x + t)) \ borel_measurable (lebesgue_on {-pi..pi})" using absolutely_integrable_on_def fxt integrable_imp_measurable by blast show "(\x. sin (x/2)) \ borel_measurable (lebesgue_on {-pi..pi})" by (intro continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto qed auto have "(\x. f(x + t) - f t) absolutely_integrable_on {-pi..pi}" by (intro set_integral_diff fxt) (simp add: set_integrable_def) moreover have "(\x. \x\ powr (a - 1)) absolutely_integrable_on {-pi..pi}" proof - have "((\x. x powr (a - 1)) has_integral inverse a * pi powr a - inverse a * 0 powr a) {0..pi}" proof (rule fundamental_theorem_of_calculus_interior) show "continuous_on {0..pi} (\x. inverse a * x powr a)" using \a > 0\ by (intro continuous_on_powr' continuous_intros) auto show "((\b. inverse a * b powr a) has_vector_derivative x powr (a - 1)) (at x)" if "x \ {0<..a > 0\ apply (simp flip: has_field_derivative_iff_has_vector_derivative) apply (rule derivative_eq_intros | simp)+ done qed auto then have int: "(\x. x powr (a - 1)) integrable_on {0..pi}" by blast show ?thesis proof (rule nonnegative_absolutely_integrable_1) show "(\x. \x\ powr (a - 1)) integrable_on {-pi..pi}" proof (rule Henstock_Kurzweil_Integration.integrable_combine) show "(\x. \x\ powr (a - 1)) integrable_on {0..pi}" using int integrable_eq by fastforce then show "(\x. \x\ powr (a - 1)) integrable_on {- pi..0}" using Henstock_Kurzweil_Integration.integrable_reflect_real by fastforce qed auto qed auto qed ultimately show "?g integrable_on {-pi..pi}" apply (intro set_lebesgue_integral_eq_integral absolutely_integrable_max_1 absolutely_integrable_bounded_measurable_product_real set_integrable_abs measurable) apply (auto simp: image_constant_conv) done show "norm ((f(x + t) - f t) / sin (x/2)) \ ?g x" if "x \ {-pi..pi}" for x proof (cases "\x\ < e") case True then show ?thesis using e [OF True] by (simp add: max_def) next case False then have "\inverse (sin (x/2))\ \ B" using B that by force then have "\inverse (sin (x/2))\ * \f(x + t) - f t\ \ B * \f(x + t) - f t\" by (simp add: mult_right_mono) then have "\f(x + t) - f t\ / \sin (x/2)\ \ B * \f(x + t) - f t\" by (simp add: divide_simps split: if_split_asm) then show ?thesis by auto qed qed auto qed (auto simp: periodic) text\In particular, a Lipschitz condition at the point\ corollary Lipschitz_Fourier_convergence_periodic: assumes f: "f absolutely_integrable_on {-pi..pi}" and "d > 0" and ft: "\x. \x-t\ < d \ \f x - f t\ \ M * \x-t\" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ f t" using Hoelder_Fourier_convergence_periodic [OF f \d > 0\, of 1] assms by (fastforce simp add:) text\In particular, if left and right derivatives both exist\ proposition bi_differentiable_Fourier_convergence_periodic: assumes f: "f absolutely_integrable_on {-pi..pi}" and f_lt: "f differentiable at_left t" and f_gt: "f differentiable at_right t" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ f t" proof - obtain D_lt where D_lt: "\e. e > 0 \ \d>0. \x dist x t < d \ dist ((f x - f t) / (x - t)) D_lt < e" using f_lt unfolding has_field_derivative_iff real_differentiable_def Lim_within by (meson lessThan_iff) then obtain d_lt where "d_lt > 0" and d_lt: "\x. \x < t; 0 < \x - t\; \x - t\ < d_lt\ \ \(f x - f t) / (x - t) - D_lt\ < 1" unfolding dist_real_def by (metis dist_commute dist_real_def zero_less_one) obtain D_gt where D_gt: "\e. e > 0 \ \d>0. \x>t. 0 < dist x t \ dist x t < d \ dist ((f x - f t) / (x - t)) D_gt < e" using f_gt unfolding has_field_derivative_iff real_differentiable_def Lim_within by (meson greaterThan_iff) then obtain d_gt where "d_gt > 0" and d_gt: "\x. \x > t; 0 < \x - t\; \x - t\ < d_gt\ \ \(f x - f t) / (x - t) - D_gt\ < 1" unfolding dist_real_def by (metis dist_commute dist_real_def zero_less_one) show ?thesis proof (rule Lipschitz_Fourier_convergence_periodic [OF f]) fix x assume "\x - t\ < min d_lt d_gt" then have *: "\x - t\ < d_lt" "\x - t\ < d_gt" by auto consider "xt" by linarith then show "\f x - f t\ \ (\D_lt\ + \D_gt\ + 1) * \x - t\" proof cases case 1 then have "\(f x - f t) / (x - t) - D_lt\ < 1" using d_lt [OF 1] * by auto then have "\(f x - f t) / (x - t)\ - \D_lt\ < 1" by linarith then have "\f x - f t\ \ (\D_lt\ + 1) * \x - t\" by (simp add: comm_semiring_class.distrib divide_simps split: if_split_asm) also have "\ \ (\D_lt\ + \D_gt\ + 1) * \x - t\" by (simp add: mult_right_mono) finally show ?thesis . next case 3 then have "\(f x - f t) / (x - t) - D_gt\ < 1" using d_gt [OF 3] * by auto then have "\(f x - f t) / (x - t)\ - \D_gt\ < 1" by linarith then have "\f x - f t\ \ (\D_gt\ + 1) * \x - t\" by (simp add: comm_semiring_class.distrib divide_simps split: if_split_asm) also have "\ \ (\D_lt\ + \D_gt\ + 1) * \x - t\" by (simp add: mult_right_mono) finally show ?thesis . qed auto qed (auto simp: \0 < d_gt\ \0 < d_lt\ periodic) qed text\And in particular at points where the function is differentiable\ lemma differentiable_Fourier_convergence_periodic: assumes f: "f absolutely_integrable_on {-pi..pi}" and fdif: "f differentiable (at t)" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ f t" by (rule bi_differentiable_Fourier_convergence_periodic) (auto simp: differentiable_at_withinI assms) text\Use reflection to halve the region of integration\ lemma absolutely_integrable_mult_Dirichlet_kernel_reflected: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\x. Dirichlet_kernel n x * f(t+x)) absolutely_integrable_on {-pi..pi}" "(\x. Dirichlet_kernel n x * f(t-x)) absolutely_integrable_on {-pi..pi}" "(\x. Dirichlet_kernel n x * c) absolutely_integrable_on {-pi..pi}" proof - show "(\x. Dirichlet_kernel n x * f(t+x)) absolutely_integrable_on {-pi..pi}" apply (rule absolutely_integrable_mult_Dirichlet_kernel) using absolutely_integrable_periodic_offset [OF f, of t] periodic apply simp done then show "(\x. Dirichlet_kernel n x * f(t-x)) absolutely_integrable_on {-pi..pi}" by (subst absolutely_integrable_reflect_real [symmetric]) simp show "(\x. Dirichlet_kernel n x * c) absolutely_integrable_on {-pi..pi}" by (simp add: absolutely_integrable_measurable_real borel_measurable_integrable integrable_Dirichlet_kernel) qed lemma absolutely_integrable_mult_Dirichlet_kernel_reflected_part: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "d \ pi" shows "(\x. Dirichlet_kernel n x * f(t+x)) absolutely_integrable_on {0..d}" "(\x. Dirichlet_kernel n x * f(t-x)) absolutely_integrable_on {0..d}" "(\x. Dirichlet_kernel n x * c) absolutely_integrable_on {0..d}" using absolutely_integrable_mult_Dirichlet_kernel_reflected [OF f periodic, of n] \d \ pi\ by (force intro: absolutely_integrable_on_subinterval)+ lemma absolutely_integrable_mult_Dirichlet_kernel_reflected_part2: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "d \ pi" shows "(\x. Dirichlet_kernel n x * (f(t+x) + f(t-x))) absolutely_integrable_on {0..d}" "(\x. Dirichlet_kernel n x * ((f(t+x) + f(t-x)) - c)) absolutely_integrable_on {0..d}" using absolutely_integrable_mult_Dirichlet_kernel_reflected_part assms by (simp_all add: distrib_left right_diff_distrib) lemma integral_reflect_and_add: fixes f :: "real \ 'b::euclidean_space" assumes "integrable (lebesgue_on {-a..a}) f" shows "integral\<^sup>L (lebesgue_on {-a..a}) f = integral\<^sup>L (lebesgue_on {0..a}) (\x. f x + f(-x))" proof (cases "a \ 0") case True have f: "integrable (lebesgue_on {0..a}) f" "integrable (lebesgue_on {-a..0}) f" using integrable_subinterval assms by fastforce+ then have fm: "integrable (lebesgue_on {0..a}) (\x. f(-x))" using integrable_reflect_real by fastforce have "integral\<^sup>L (lebesgue_on {-a..a}) f = integral\<^sup>L (lebesgue_on {-a..0}) f + integral\<^sup>L (lebesgue_on {0..a}) f" by (simp add: True assms integral_combine) also have "\ = integral\<^sup>L (lebesgue_on {0..a}) (\x. f(-x)) + integral\<^sup>L (lebesgue_on {0..a}) f" by (metis (no_types) add.inverse_inverse add.inverse_neutral integral_reflect_real) also have "\ = integral\<^sup>L (lebesgue_on {0..a}) (\x. f x + f(-x))" by (simp add: fm f) finally show ?thesis . qed (use assms in auto) lemma Fourier_sum_offset_Dirichlet_kernel_half: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\k\2*n. Fourier_coefficient f k * trigonometric_set k t) - l = (LINT x|lebesgue_on {0..pi}. Dirichlet_kernel n x * (f(t+x) + f(t-x) - 2*l)) / pi" proof - have fxt: "(\x. f(x + t)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset assms by auto have int: "integrable (lebesgue_on {0..pi}) (Dirichlet_kernel n)" using not_integrable_integral_eq by fastforce have "LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * f(x + t) = LINT x|lebesgue_on {0..pi}. Dirichlet_kernel n x * f(x + t) + Dirichlet_kernel n (- x) * f(- x + t)" by (simp add: integral_reflect_and_add absolutely_integrable_imp_integrable absolutely_integrable_mult_Dirichlet_kernel fxt) also have "\ = (LINT x|lebesgue_on {0..pi}. Dirichlet_kernel n x * (f(t+x) + f(t-x) - 2*l)) + pi * l" using absolutely_integrable_mult_Dirichlet_kernel_reflected_part [OF f periodic order_refl, of n t] apply (simp add: algebra_simps absolutely_integrable_imp_integrable int) done finally show ?thesis by (simp add: Fourier_sum_offset_Dirichlet_kernel assms divide_simps) qed lemma Fourier_sum_limit_Dirichlet_kernel_half: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ l \ (\n. (LINT x|lebesgue_on {0..pi}. Dirichlet_kernel n x * (f(t+x) + f(t-x) - 2*l))) \ 0" apply (simp flip: Fourier_sum_limit_pair [OF f]) apply (simp add: Lim_null [where l=l] Fourier_sum_offset_Dirichlet_kernel_half assms) done subsection\Localization principle: convergence only depends on values "nearby"\ proposition Riemann_localization_integral: assumes f: "f absolutely_integrable_on {-pi..pi}" and g: "g absolutely_integrable_on {-pi..pi}" and "d > 0" and d: "\x. \x\ < d \ f x = g x" shows "(\n. integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. Dirichlet_kernel n x * f x) - integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. Dirichlet_kernel n x * g x)) \ 0" (is "?a \ 0") proof - let ?f = "\x. (if \x\ < d then 0 else f x - g x) / sin(x/2) / 2" have eq: "?a n = integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin((n+1/2) * x) * ?f x)" for n proof - have "?a n = integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. Dirichlet_kernel n x * (if \x\ < d then 0 else f x - g x))" apply (simp add: absolutely_integrable_imp_integrable absolutely_integrable_mult_Dirichlet_kernel f g flip: Bochner_Integration.integral_diff) apply (auto simp: d algebra_simps intro: Bochner_Integration.integral_cong) done also have "\ = integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin((n+1/2) * x) * ?f x)" using \d > 0\ by (auto simp: Dirichlet_kernel_def intro: Bochner_Integration.integral_cong) finally show ?thesis . qed show ?thesis unfolding eq proof (rule Riemann_lebesgue_sin_half) obtain B where "B>0" and B: "\x. x \ ({-pi..pi} - {-d<.. \inverse (sin (x/2))\ \ B" using bounded_inverse_sin_half [OF \d > 0\] by metis have "(\x. (if \x\ < d then 0 else f x - g x) / sin (x/2)) absolutely_integrable_on {-pi..pi}" proof (rule measurable_bounded_by_integrable_imp_absolutely_integrable) show "(\x. (if \x\ < d then 0 else f x - g x) / sin (x/2)) \ borel_measurable (lebesgue_on {-pi..pi})" proof (intro measurable) show "f \ borel_measurable (lebesgue_on {-pi..pi})" "g \ borel_measurable (lebesgue_on {-pi..pi})" using absolutely_integrable_on_def f g integrable_imp_measurable by blast+ show "(\x. x) \ borel_measurable (lebesgue_on {-pi..pi})" "(\x. sin (x/2)) \ borel_measurable (lebesgue_on {-pi..pi})" by (intro continuous_intros continuous_imp_measurable_on_sets_lebesgue | force)+ qed auto have "(\x. B * \f x - g x\) absolutely_integrable_on {-pi..pi}" using \B > 0\ by (simp add: f g set_integrable_abs) then show "(\x. B * \f x - g x\) integrable_on {-pi..pi}" using absolutely_integrable_on_def by blast next fix x assume x: "x \ {-pi..pi}" then have [simp]: "sin (x/2) = 0 \ x=0" using \0 < d\ by (force simp: sin_zero_iff) show "norm ((if \x\ < d then 0 else f x - g x) / sin (x/2)) \ B * \f x - g x\" proof (cases "\x\ < d") case False then have "1 \ B * \sin (x/2)\" using \B > 0\ B [of x] x \d > 0\ by (auto simp: abs_less_iff divide_simps split: if_split_asm) then have "1 * \f x - g x\ \ B * \sin (x/2)\ * \f x - g x\" by (metis (full_types) abs_ge_zero mult.commute mult_left_mono) then show ?thesis using False by (auto simp: divide_simps algebra_simps) qed (simp add: d) qed auto then show "(\x. (if \x\ < d then 0 else f x - g x) / sin (x/2) / 2) absolutely_integrable_on {-pi..pi}" using set_integrable_divide by blast qed qed lemma Riemann_localization_integral_range: assumes f: "f absolutely_integrable_on {-pi..pi}" and "0 < d" "d \ pi" shows "(\n. integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. Dirichlet_kernel n x * f x) - integral\<^sup>L (lebesgue_on {-d..d}) (\x. Dirichlet_kernel n x * f x)) \ 0" proof - have *: "(\n. (LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * f x) - (LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * (if x \ {-d..d} then f x else 0))) \ 0" proof (rule Riemann_localization_integral [OF f _ \0 < d\]) have "f absolutely_integrable_on {-d..d}" using f absolutely_integrable_on_subinterval \d \ pi\ by fastforce moreover have "(\x. if - pi \ x \ x \ pi then if x \ {-d..d} then f x else 0 else 0) = (\x. if x \ {-d..d} then f x else 0)" using \d \ pi\ by force ultimately show "(\x. if x \ {-d..d} then f x else 0) absolutely_integrable_on {-pi..pi}" apply (subst absolutely_integrable_restrict_UNIV [symmetric]) apply (simp flip: absolutely_integrable_restrict_UNIV [of "{-d..d}" f]) done qed auto then show ?thesis using integral_restrict [of "{-d..d}" "{-pi..pi}" "\x. Dirichlet_kernel _ x * f x"] assms by (simp add: if_distrib cong: if_cong) qed lemma Riemann_localization: assumes f: "f absolutely_integrable_on {-pi..pi}" and g: "g absolutely_integrable_on {-pi..pi}" and perf: "\x. f(x + 2*pi) = f x" and perg: "\x. g(x + 2*pi) = g x" and "d > 0" and d: "\x. \x-t\ < d \ f x = g x" shows "(\n. \k\n. Fourier_coefficient f k * trigonometric_set k t) \ c \ (\n. \k\n. Fourier_coefficient g k * trigonometric_set k t) \ c" proof - have "(\n. LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * f(x + t)) \ pi * c \ (\n. LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * g (x + t)) \ pi * c" proof (intro Lim_transform_eq Riemann_localization_integral) show "(\x. f(x + t)) absolutely_integrable_on {-pi..pi}" "(\x. g (x + t)) absolutely_integrable_on {-pi..pi}" using assms by (auto intro: absolutely_integrable_periodic_offset) qed (use assms in auto) then show ?thesis by (simp add: Fourier_sum_limit_Dirichlet_kernel assms) qed subsection\Localize the earlier integral\ lemma Riemann_localization_integral_range_half: assumes f: "f absolutely_integrable_on {-pi..pi}" and "0 < d" "d \ pi" shows "(\n. (LINT x|lebesgue_on {0..pi}. Dirichlet_kernel n x * (f x + f(-x))) - (LINT x|lebesgue_on {0..d}. Dirichlet_kernel n x * (f x + f(-x)))) \ 0" proof - have *: "(\x. Dirichlet_kernel n x * f x) absolutely_integrable_on {-d..d}" for n by (metis (full_types) absolutely_integrable_mult_Dirichlet_kernel absolutely_integrable_on_subinterval \d \ pi\ atLeastatMost_subset_iff f neg_le_iff_le) show ?thesis using absolutely_integrable_mult_Dirichlet_kernel [OF f] using Riemann_localization_integral_range [OF assms] apply (simp add: "*" absolutely_integrable_imp_integrable integral_reflect_and_add) apply (simp add: algebra_simps) done qed lemma Fourier_sum_limit_Dirichlet_kernel_part: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and d: "0 < d" "d \ pi" shows "(\n. \k\n. Fourier_coefficient f k * trigonometric_set k t) \ l \ (\n. (LINT x|lebesgue_on {0..d}. Dirichlet_kernel n x * ((f(t+x) + f(t-x)) - 2*l))) \ 0" proof - have "(\x. f(t+x)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset [OF f, of t] periodic by simp then have int: "(\x. f(t+x) - l) absolutely_integrable_on {-pi..pi}" by auto have "(\n. LINT x|lebesgue_on {0..pi}. Dirichlet_kernel n x * (f(t+x) + f(t-x) - 2*l)) \ 0 \ (\n. LINT x|lebesgue_on {0..d}. Dirichlet_kernel n x * (f(t+x) + f(t-x) - 2*l)) \ 0" by (rule Lim_transform_eq) (use Riemann_localization_integral_range_half [OF int d] in auto) then show ?thesis by (simp add: Fourier_sum_limit_Dirichlet_kernel_half assms) qed subsection\Make a harmless simplifying tweak to the Dirichlet kernel\ lemma inte_Dirichlet_kernel_mul_expand: assumes f: "f \ borel_measurable (lebesgue_on S)" and S: "S \ sets lebesgue" shows "(LINT x|lebesgue_on S. Dirichlet_kernel n x * f x = LINT x|lebesgue_on S. sin((n+1/2) * x) * f x / (2 * sin(x/2))) \ (integrable (lebesgue_on S) (\x. Dirichlet_kernel n x * f x) \ integrable (lebesgue_on S) (\x. sin((n+1/2) * x) * f x / (2 * sin(x/2))))" proof (cases "0 \ S") case True have *: "{x. x = 0 \ x \ S} \ sets (lebesgue_on S)" using True by (simp add: S sets_restrict_space_iff cong: conj_cong) have bm1: "(\x. Dirichlet_kernel n x * f x) \ borel_measurable (lebesgue_on S)" unfolding Dirichlet_kernel_def by (force intro: * assms measurable continuous_imp_measurable_on_sets_lebesgue continuous_intros) have bm2: "(\x. sin ((n + 1/2) * x) * f x / (2 * sin (x/2))) \ borel_measurable (lebesgue_on S)" by (intro assms measurable continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto have 0: "{0} \ null_sets (lebesgue_on S)" using True by (simp add: S null_sets_restrict_space) show ?thesis apply (intro conjI integral_cong_AE integrable_cong_AE bm1 bm2 AE_I' [OF 0]) unfolding Dirichlet_kernel_def by auto next case False then show ?thesis unfolding Dirichlet_kernel_def by (auto intro!: Bochner_Integration.integral_cong Bochner_Integration.integrable_cong) qed lemma assumes f: "f \ borel_measurable (lebesgue_on S)" and S: "S \ sets lebesgue" shows integral_Dirichlet_kernel_mul_expand: "(LINT x|lebesgue_on S. Dirichlet_kernel n x * f x) = (LINT x|lebesgue_on S. sin((n+1/2) * x) * f x / (2 * sin(x/2)))" (is "?th1") and integrable_Dirichlet_kernel_mul_expand: "integrable (lebesgue_on S) (\x. Dirichlet_kernel n x * f x) \ integrable (lebesgue_on S) (\x. sin((n+1/2) * x) * f x / (2 * sin(x/2)))" (is "?th2") using inte_Dirichlet_kernel_mul_expand [OF assms] by auto proposition Fourier_sum_limit_sine_part: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and d: "0 < d" "d \ pi" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ l \ (\n. LINT x|lebesgue_on {0..d}. sin((n + 1/2) * x) * ((f(t+x) + f(t-x) - 2*l) / x)) \ 0" (is "?lhs \ ?\ \ 0") proof - have ftx: "(\x. f(t+x)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset assms by auto then have ftmx: "(\x. f(t-x)) absolutely_integrable_on {-pi..pi}" by (simp flip: absolutely_integrable_reflect_real [where f= "(\x. f(t+x))"]) have fbm: "(\x. f(t+x) + f(t-x) - 2*l) absolutely_integrable_on {-pi..pi}" by (force intro: ftmx ftx) let ?\ = "\n. LINT x|lebesgue_on {0..d}. Dirichlet_kernel n x * ((f(t+x) + f(t-x)) - 2*l)" have "?lhs \ ?\ \ 0" by (intro Fourier_sum_limit_Dirichlet_kernel_part assms) moreover have "?\ \ 0 \ ?\ \ 0" proof (rule Lim_transform_eq [OF Lim_transform_eventually]) let ?f = "\n. LINT x|lebesgue_on {0..d}. sin((real n + 1/2) * x) * (1 / (2 * sin(x/2)) - 1/x) * (f(t+x) + f(t-x) - 2*l)" have "?f n = (LINT x|lebesgue_on {-pi..pi}. sin ((n + 1/2) * x) * ((if x \ {0..d} then 1 / (2 * sin (x/2)) - 1/x else 0) * (f(t+x) + f(t-x) - 2*l)))" for n using d by (simp add: integral_restrict if_distrib [of "\u. _ * (u * _)"] mult.assoc flip: atLeastAtMost_iff cong: if_cong) moreover have "\ \ 0" proof (intro Riemann_lebesgue_sin_half absolutely_integrable_bounded_measurable_product_real) have "(\x. 1 / (2 * sin(x/2)) - 1/x) \ borel_measurable (lebesgue_on {0..d})" by (intro measurable continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto then show "(\x. if x \ {0..d} then 1 / (2 * sin(x/2)) - 1/x else 0) \ borel_measurable (lebesgue_on {-pi..pi})" using d by (simp add: borel_measurable_if_lebesgue_on flip: atLeastAtMost_iff) let ?g = "\x::real. 1 / (2 * sin(x/2)) - 1/x" have limg: "(?g \ ?g a) (at a within {0..d})" \\thanks to Manuel Eberl\ if a: "0 \ a" "a \ d" for a proof - have "(?g \ 0) (at_right 0)" and "(?g \ ?g d) (at_left d)" using d sin_gt_zero[of "d/2"] by (real_asymp simp: field_simps)+ moreover have "(?g \ ?g a) (at a)" if "a > 0" using a that d sin_gt_zero[of "a/2"] that by (real_asymp simp: field_simps) ultimately show ?thesis using that d by (cases "a = 0 \ a = d") (auto simp: at_within_Icc_at at_within_Icc_at_right at_within_Icc_at_left) qed have "((\x. if x \ {0..d} then 1 / (2 * sin(x/2)) - 1/x else 0) ` {-pi..pi}) = ((\x. 1 / (2 * sin(x/2)) - 1/x) ` {0..d})" using d by (auto intro: image_eqI [where x=0]) moreover have "bounded \" by (intro compact_imp_bounded compact_continuous_image) (auto simp: continuous_on limg) ultimately show "bounded ((\x. if x \ {0..d} then 1 / (2 * sin(x/2)) - 1/x else 0) ` {-pi..pi})" by simp qed (auto simp: fbm) ultimately show "?f \ (0::real)" by simp show "\\<^sub>F n in sequentially. ?f n = ?\ n - ?\ n" proof (rule eventually_sequentiallyI [where c = "Suc 0"]) fix n assume "n \ Suc 0" have "integrable (lebesgue_on {0..d}) (\x. sin ((real n + 1/2) * x) * (f(t+x) + f(t-x) - 2*l) / (2 * sin(x/2)))" using d apply (subst integrable_Dirichlet_kernel_mul_expand [symmetric]) apply (intro assms absolutely_integrable_imp_borel_measurable absolutely_integrable_on_subinterval [OF fbm] absolutely_integrable_imp_integrable absolutely_integrable_mult_Dirichlet_kernel_reflected_part2 | force)+ done moreover have "integrable (lebesgue_on {0..d}) (\x. sin ((real n + 1/2) * x) * (f(t+x) + f(t-x) - 2*l) / x)" proof (rule integrable_cong_AE_imp) let ?g = "\x. Dirichlet_kernel n x * (2 * sin(x/2) / x * (f(t+x) + f(t-x) - 2*l))" have *: "\2 * sin (x/2) / x\ \ 1" for x::real using abs_sin_x_le_abs_x [of "x/2"] by (simp add: divide_simps mult.commute abs_mult) have "bounded ((\x. 1 + (x/2)\<^sup>2) ` {-pi..pi})" by (intro compact_imp_bounded compact_continuous_image continuous_intros) auto then have bo: "bounded ((\x. 2 * sin (x/2) / x) ` {-pi..pi})" using * unfolding bounded_real by blast show "integrable (lebesgue_on {0..d}) ?g" using d apply (intro absolutely_integrable_imp_integrable absolutely_integrable_on_subinterval [OF absolutely_integrable_mult_Dirichlet_kernel] absolutely_integrable_bounded_measurable_product_real bo fbm measurable continuous_imp_measurable_on_sets_lebesgue continuous_intros, auto) done show "(\x. sin ((n + 1/2) * x) * (f(t+x) + f(t-x) - 2*l) / x) \ borel_measurable (lebesgue_on {0..d})" using d apply (intro measurable absolutely_integrable_imp_borel_measurable absolutely_integrable_on_subinterval [OF ftx] absolutely_integrable_on_subinterval [OF ftmx] absolutely_integrable_continuous_real continuous_intros, auto) done have "Dirichlet_kernel n x * (2 * sin(x/2)) / x = sin ((real n + 1/2) * x) / x" if "0 < x" "x \ d" for x using that d by (simp add: Dirichlet_kernel_def divide_simps sin_zero_pi_iff) then show "AE x in lebesgue_on {0..d}. ?g x = sin ((real n + 1/2) * x) * (f(t+x) + f(t-x) - 2*l) / x" using d by (force intro: AE_I' [where N="{0}"]) qed ultimately have "?f n = (LINT x|lebesgue_on {0..d}. sin ((n + 1/2) * x) * (f(t+x) + f(t-x) - 2*l) / (2 * sin(x/2))) - (LINT x|lebesgue_on {0..d}. sin ((n + 1/2) * x) * (f(t+x) + f(t-x) - 2*l) / x)" by (simp add: right_diff_distrib [of _ _ "1/_"] left_diff_distrib) also have "\ = ?\ n - ?\ n" using d by (simp add: measurable_restrict_mono [OF absolutely_integrable_imp_borel_measurable [OF fbm]] integral_Dirichlet_kernel_mul_expand) finally show "?f n = ?\ n - ?\ n" . qed qed ultimately show ?thesis by simp qed subsection\Dini's test for the convergence of a Fourier series\ proposition Fourier_Dini_test: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and int0d: "integrable (lebesgue_on {0..d}) (\x. \f(t+x) + f(t-x) - 2*l\ / x)" and "0 < d" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ l" proof - define \ where "\ \ \n x. sin((n + 1/2) * x) * ((f(t+x) + f(t-x) - 2*l) / x)" have "(\n. LINT x|lebesgue_on {0..pi}. \ n x) \ 0" unfolding lim_sequentially dist_real_def proof (intro allI impI) fix e :: real assume "e > 0" define \ where "\ \ \x. LINT x|lebesgue_on {0..x}. \f(t+x) + f(t-x) - 2*l\ / x" have [simp]: "\ 0 = 0" by (simp add: \_def integral_eq_zero_null_sets) have cont: "continuous_on {0..d} \" unfolding \_def using indefinite_integral_continuous_real int0d by blast with \d > 0\ have "\e>0. \da>0. \x'\{0..d}. \x' - 0\ < da \ \\ x' - \ 0\ < e" by (force simp: continuous_on_real_range dest: bspec [where x=0]) with \e > 0\ obtain k where "k > 0" and k: "\x'. \0 \ x'; x' < k; x' \ d\ \ \\ x'\ < e/2" by (drule_tac x="e/2" in spec) auto define \ where "\ \ min d (min (k/2) pi)" have e2: "\\ \\ < e/2" and "\ > 0" "\ \ pi" unfolding \_def using k \k > 0\ \d > 0\ by auto have ftx: "(\x. f(t+x)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset assms by auto then have ftmx: "(\x. f(t-x)) absolutely_integrable_on {-pi..pi}" by (simp flip: absolutely_integrable_reflect_real [where f= "(\x. f(t+x))"]) have 1: "\ n absolutely_integrable_on {0..\}" for n::nat unfolding \_def proof (rule absolutely_integrable_bounded_measurable_product_real) show "(\x. sin ((real n + 1/2) * x)) \ borel_measurable (lebesgue_on {0..\})" by (intro continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto show "bounded ((\x. sin ((real n + 1/2) * x)) ` {0..\})" by (simp add: bounded_iff) (use abs_sin_le_one in blast) have "(\x. (f(t+x) + f(t-x) - 2*l) / x) \ borel_measurable (lebesgue_on {0..\})" using \d > 0\ unfolding \_def by (intro measurable absolutely_integrable_imp_borel_measurable absolutely_integrable_on_subinterval [OF ftx] absolutely_integrable_on_subinterval [OF ftmx] absolutely_integrable_continuous_real continuous_intros, auto) moreover have "integrable (lebesgue_on {0..\}) (norm \ (\x. (f(t+x) + f(t-x) - 2*l) / x))" proof (rule integrable_subinterval [of 0 d]) show "integrable (lebesgue_on {0..d}) (norm \ (\x. (f(t+x) + f(t-x) - 2*l) / x))" using int0d by (subst integrable_cong) (auto simp: o_def) show "{0..\} \ {0..d}" using \d > 0\ by (auto simp: \_def) qed ultimately show "(\x. (f(t+x) + f(t-x) - 2*l) / x) absolutely_integrable_on {0..\}" by (auto simp: absolutely_integrable_measurable) qed auto have 2: "\ n absolutely_integrable_on {\..pi}" for n::nat unfolding \_def proof (rule absolutely_integrable_bounded_measurable_product_real) show "(\x. sin ((n + 1/2) * x)) \ borel_measurable (lebesgue_on {\..pi})" by (intro continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto show "bounded ((\x. sin ((n + 1/2) * x)) ` {\..pi})" by (simp add: bounded_iff) (use abs_sin_le_one in blast) have "(\x. inverse x * (f(t+x) + f(t-x) - 2*l)) absolutely_integrable_on {\..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) have "(\x. 1/x) \ borel_measurable (lebesgue_on {\..pi})" by (auto simp: measurable_completion measurable_restrict_space1) then show "inverse \ borel_measurable (lebesgue_on {\..pi})" by (metis (no_types, lifting) inverse_eq_divide measurable_lebesgue_cong) have "\x\{\..pi}. inverse \x\ \ inverse \" using \0 < \\ by auto then show "bounded (inverse ` {\..pi})" by (auto simp: bounded_iff) show "(\x. f(t+x) + f(t-x) - 2*l) absolutely_integrable_on {\..pi}" proof (rule absolutely_integrable_on_subinterval) show "(\x. (f(t+x) + f(t-x) - 2*l)) absolutely_integrable_on {-pi..pi}" by (fast intro: ftx ftmx absolutely_integrable_on_const) show "{\..pi} \ {-pi..pi}" using \0 < \\ by auto qed qed auto then show "(\x. (f(t+x) + f(t-x) - 2*l) / x) absolutely_integrable_on {\..pi}" by (metis (no_types, lifting) divide_inverse mult.commute set_integrable_cong) qed auto have "(\x. f(t+x) - l) absolutely_integrable_on {-pi..pi}" using ftx by auto moreover have "bounded ((\x. 0) ` {x. \x\ < \})" using bounded_def by blast moreover have "bounded (inverse ` {x. \ \x\ < \})" using \\ > 0\ by (auto simp: divide_simps intro: boundedI [where B = "1/\"]) ultimately have "(\x. (if \x\ < \ then 0 else inverse x) * (if x \ {-pi..pi} then f(t+x) - l else 0)) absolutely_integrable_on UNIV" apply (intro absolutely_integrable_bounded_measurable_product_real measurable set_integral_diff) apply (auto simp: lebesgue_on_UNIV_eq measurable_completion simp flip: absolutely_integrable_restrict_UNIV [where S = "{-pi..pi}"]) done moreover have "(if x \ {-pi..pi} then if \x\ < \ then 0 else (f(t+x) - l) / x else 0) = (if \x\ < \ then 0 else inverse x) * (if x \ {-pi..pi} then f(t+x) - l else 0)" for x by (auto simp: divide_simps) ultimately have *: "(\x. if \x\ < \ then 0 else (f(t+x) - l) / x) absolutely_integrable_on {-pi..pi}" by (simp add: flip: absolutely_integrable_restrict_UNIV [where S = "{-pi..pi}"]) have "(\n. LINT x|lebesgue_on {0..pi}. sin ((n + 1/2) * x) * (if \x\ < \ then 0 else (f(t+x) - l) / x) + sin ((n + 1/2) * -x) * (if \x\ < \ then 0 else (f(t-x) - l) / -x)) \ 0" using Riemann_lebesgue_sin_half [OF *] * by (simp add: integral_reflect_and_add absolutely_integrable_imp_integrable absolutely_integrable_sin_product cong: if_cong) moreover have "sin ((n + 1/2) * x) * (if \x\ < \ then 0 else (f(t+x) - l) / x) + sin ((n + 1/2) * -x) * (if \x\ < \ then 0 else (f(t-x) - l) / -x) = (if \ \x\ < \ then \ n x else 0)" for x n by simp (auto simp: divide_simps algebra_simps \_def) ultimately have "(\n. LINT x|lebesgue_on {0..pi}. (if \ \x\ < \ then \ n x else 0)) \ 0" by simp moreover have "(if 0 \ x \ x \ pi then if \ \x\ < \ then \ n x else 0 else 0) = (if \ \ x \ x \ pi then \ n x else 0)" for x n using \\ > 0\ by auto ultimately have \: "(\n. LINT x|lebesgue_on {\..pi}. \ n x) \ 0" by (simp flip: Lebesgue_Measure.integral_restrict_UNIV) then obtain N::nat where N: "\n. n\N \ \LINT x|lebesgue_on {\..pi}. \ n x\ < e/2" unfolding lim_sequentially dist_real_def by (metis (full_types) \0 < e\ diff_zero half_gt_zero_iff) have "\integral\<^sup>L (lebesgue_on {0..pi}) (\ n)\ < e" if "n \ N" for n::nat proof - have int: "integrable (lebesgue_on {0..pi}) (\ (real n))" by (intro integrable_combine [of concl: 0 pi] absolutely_integrable_imp_integrable) (use \\ > 0\ \\ \ pi\ 1 2 in auto) then have "integral\<^sup>L (lebesgue_on {0..pi}) (\ n) = integral\<^sup>L (lebesgue_on {0..\}) (\ n) + integral\<^sup>L (lebesgue_on {\..pi}) (\ n)" by (rule integral_combine) (use \0 < \\ \\ \ pi\ in auto) moreover have "\integral\<^sup>L (lebesgue_on {0..\}) (\ n)\ \ LINT x|lebesgue_on {0..\}. \f(t + x) + f(t - x) - 2 * l\ / x" proof (rule integral_abs_bound_integral) show "integrable (lebesgue_on {0..\}) (\ (real n))" by (meson integrable_subinterval \\ \ pi\ int atLeastatMost_subset_iff order_refl) have "{0..\} \ {0..d}" by (auto simp: \_def) then show "integrable (lebesgue_on {0..\}) (\x. \f(t + x) + f(t - x) - 2 * l\ / x)" by (rule integrable_subinterval [OF int0d]) show "\\ (real n) x\ \ \f(t + x) + f(t - x) - 2 * l\ / x" if "x \ space (lebesgue_on {0..\})" for x using that apply (auto simp: \_def divide_simps abs_mult) by (simp add: mult.commute mult_left_le) qed ultimately have "\integral\<^sup>L (lebesgue_on {0..pi}) (\ n)\ < e/2 + e/2" using N [OF that] e2 unfolding \_def by linarith then show ?thesis by simp qed then show "\N. \n\N. \integral\<^sup>L (lebesgue_on {0..pi}) (\ (real n)) - 0\ < e" by force qed then show ?thesis unfolding \_def using Fourier_sum_limit_sine_part assms pi_gt_zero by blast qed subsection\Cesaro summability of Fourier series using Fejér kernel\ definition Fejer_kernel :: "nat \ real \ real" where "Fejer_kernel \ \n x. if n = 0 then 0 else (\r sin(x/2) = 0") case True have "(\rrx. Fejer_kernel n x * f x) absolutely_integrable_on {-pi..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) show "Fejer_kernel n \ borel_measurable (lebesgue_on {-pi..pi})" by (simp add: Fejer_kernel_continuous continuous_imp_measurable_on_sets_lebesgue) show "bounded (Fejer_kernel n ` {-pi..pi})" using Fejer_kernel_continuous compact_continuous_image compact_imp_bounded by blast qed (use assms in auto) lemma absolutely_integrable_mult_Fejer_kernel_reflected1: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\x. Fejer_kernel n x * f(t + x)) absolutely_integrable_on {-pi..pi}" using assms by (force intro: absolutely_integrable_mult_Fejer_kernel absolutely_integrable_periodic_offset) lemma absolutely_integrable_mult_Fejer_kernel_reflected2: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\x. Fejer_kernel n x * f(t - x)) absolutely_integrable_on {-pi..pi}" proof - have "(\x. f(t - x)) absolutely_integrable_on {-pi..pi}" using assms apply (subst absolutely_integrable_reflect_real [symmetric]) apply (simp add: absolutely_integrable_periodic_offset) done then show ?thesis by (rule absolutely_integrable_mult_Fejer_kernel) qed lemma absolutely_integrable_mult_Fejer_kernel_reflected3: shows "(\x. Fejer_kernel n x * c) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_on_const absolutely_integrable_mult_Fejer_kernel by blast lemma absolutely_integrable_mult_Fejer_kernel_reflected_part1: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "d \ pi" shows "(\x. Fejer_kernel n x * f(t + x)) absolutely_integrable_on {0..d}" by (rule absolutely_integrable_on_subinterval [OF absolutely_integrable_mult_Fejer_kernel_reflected1]) (auto simp: assms) lemma absolutely_integrable_mult_Fejer_kernel_reflected_part2: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "d \ pi" shows "(\x. Fejer_kernel n x * f(t - x)) absolutely_integrable_on {0..d}" by (rule absolutely_integrable_on_subinterval [OF absolutely_integrable_mult_Fejer_kernel_reflected2]) (auto simp: assms) lemma absolutely_integrable_mult_Fejer_kernel_reflected_part3: assumes "d \ pi" shows "(\x. Fejer_kernel n x * c) absolutely_integrable_on {0..d}" by (rule absolutely_integrable_on_subinterval [OF absolutely_integrable_mult_Fejer_kernel_reflected2]) (auto simp: assms) lemma absolutely_integrable_mult_Fejer_kernel_reflected_part4: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "d \ pi" shows "(\x. Fejer_kernel n x * (f(t + x) + f(t - x))) absolutely_integrable_on {0..d}" unfolding distrib_left by (intro set_integral_add absolutely_integrable_mult_Fejer_kernel_reflected_part1 absolutely_integrable_mult_Fejer_kernel_reflected_part2 assms) lemma absolutely_integrable_mult_Fejer_kernel_reflected_part5: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "d \ pi" shows "(\x. Fejer_kernel n x * ((f(t + x) + f(t - x)) - c)) absolutely_integrable_on {0..d}" unfolding distrib_left right_diff_distrib by (intro set_integral_add set_integral_diff absolutely_integrable_on_const absolutely_integrable_mult_Fejer_kernel_reflected_part1 absolutely_integrable_mult_Fejer_kernel_reflected_part2 assms, auto) lemma Fourier_sum_offset_Fejer_kernel_half: fixes n::nat assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "n > 0" shows "(\rk\2*r. Fourier_coefficient f k * trigonometric_set k t) / n - l = (LINT x|lebesgue_on {0..pi}. Fejer_kernel n x * (f(t + x) + f(t - x) - 2 * l)) / pi" proof - have "(\rk\2 * r. Fourier_coefficient f k * trigonometric_set k t) - real n * l = (\rk\2 * r. Fourier_coefficient f k * trigonometric_set k t) - l)" by (simp add: sum_subtractf) also have "\ = (\r = real n * ((LINT x|lebesgue_on {0..pi}. Fejer_kernel n x * (f(t + x) + f(t - x) - 2 * l)) / pi)" proof - have "integrable (lebesgue_on {0..pi}) (\x. Dirichlet_kernel i x * (f(t + x) + f(t - x) - 2 * l))" for i using absolutely_integrable_mult_Dirichlet_kernel_reflected_part2(2) f periodic by (force simp: intro!: absolutely_integrable_imp_integrable) then show ?thesis using \n > 0\ apply (simp add: Fejer_kernel_def flip: sum_divide_distrib) apply (simp add: sum_distrib_right flip: Bochner_Integration.integral_sum [symmetric]) done qed finally have "(\rk\2 * r. Fourier_coefficient f k * trigonometric_set k t) - real n * l = real n * ((LINT x|lebesgue_on {0..pi}. Fejer_kernel n x * (f(t + x) + f(t - x) - 2 * l)) / pi)" . with \n > 0\ show ?thesis by (auto simp: mult.commute divide_simps) qed lemma Fourier_sum_limit_Fejer_kernel_half: fixes n::nat assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. ((\rk\2*r. Fourier_coefficient f k * trigonometric_set k t)) / n) \ l \ ((\n. integral\<^sup>L (lebesgue_on {0..pi}) (\x. Fejer_kernel n x * ((f(t + x) + f(t - x)) - 2*l))) \ 0)" (is "?lhs = ?rhs") proof - have "?lhs \ (\n. ((\rk\2*r. Fourier_coefficient f k * trigonometric_set k t)) / n - l) \ 0" by (simp add: LIM_zero_iff) also have "\ \ (\n. ((((\rk\2*r. Fourier_coefficient f k * trigonometric_set k t)) / n) - l) * pi) \ 0" - using real_tendsto_mult_right_iff [OF pi_neq_zero] by simp + using tendsto_mult_right_iff [OF pi_neq_zero] by simp also have "\ \ ?rhs" apply (intro Lim_transform_eq [OF Lim_transform_eventually [of "\n. 0"]] eventually_sequentiallyI [of 1]) apply (simp_all add: Fourier_sum_offset_Fejer_kernel_half assms) done finally show ?thesis . qed lemma has_integral_Fejer_kernel: "has_bochner_integral (lebesgue_on {-pi..pi}) (Fejer_kernel n) (if n = 0 then 0 else pi)" proof - have "has_bochner_integral (lebesgue_on {-pi..pi}) (\x. (\rrx. (\rr 0" by (simp add: Fejer_kernel) theorem Fourier_Fejer_Cesaro_summable: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and fl: "(f \ l) (at t within atMost t)" and fr: "(f \ r) (at t within atLeast t)" shows "(\n. (\mk\2*m. Fourier_coefficient f k * trigonometric_set k t) / n) \ (l+r) / 2" proof - define h where "h \ \u. (f(t+u) + f(t-u)) - (l+r)" have "(\n. LINT u|lebesgue_on {0..pi}. Fejer_kernel n u * h u) \ 0" proof - have h0: "(h \ 0) (at 0 within atLeast 0)" proof - have l0: "((\u. f(t-u) - l) \ 0) (at 0 within {0..})" using fl unfolding Lim_within apply (elim all_forward imp_forward ex_forward conj_forward asm_rl, clarify) apply (drule_tac x="t-x" in bspec) apply (auto simp: dist_norm) done have r0: "((\u. f(t + u) - r) \ 0) (at 0 within {0..})" using fr unfolding Lim_within apply (elim all_forward imp_forward ex_forward conj_forward asm_rl, clarify) apply (drule_tac x="t+x" in bspec) apply (auto simp: dist_norm) done show ?thesis using tendsto_add [OF l0 r0] by (simp add: h_def algebra_simps) qed show ?thesis unfolding lim_sequentially dist_real_def diff_0_right proof clarify fix e::real assume "e > 0" then obtain x' where "0 < x'" "\x. \0 < x; x < x'\ \ \h x\ < e / (2 * pi)" using h0 unfolding Lim_within dist_real_def by (auto simp: dest: spec [where x="e/2/pi"]) then obtain \ where "0 < \" "\ < pi" and \: "\x. 0 < x \ x \ \ \ \h x\ < e/2/pi" apply (intro that [where \="min x' pi/2"], auto) using m2pi_less_pi by linarith have ftx: "(\x. f(t+x)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset assms by auto then have ftmx: "(\x. f(t-x)) absolutely_integrable_on {-pi..pi}" by (simp flip: absolutely_integrable_reflect_real [where f= "(\x. f(t+x))"]) have h_aint: "h absolutely_integrable_on {-pi..pi}" unfolding h_def by (intro absolutely_integrable_on_const set_integral_diff set_integral_add, auto simp: ftx ftmx) have "(\n. LINT x|lebesgue_on {\..pi}. Fejer_kernel n x * h x) \ 0" proof (rule Lim_null_comparison) define \ where "\ \ \n. (LINT x|lebesgue_on {\..pi}. \h x\ / (2 * sin(x/2) ^ 2)) / n" show "\\<^sub>F n in sequentially. norm (LINT x|lebesgue_on {\..pi}. Fejer_kernel n x * h x) \ \ n" proof (rule eventually_sequentiallyI) fix n::nat assume "n \ 1" have hint: "(\x. h x / (2 * sin(x/2) ^ 2)) absolutely_integrable_on {\..pi}" unfolding divide_inverse mult.commute [of "h _"] proof (rule absolutely_integrable_bounded_measurable_product_real) have cont: "continuous_on {\..pi} (\x. inverse (2 * (sin (x * inverse 2))\<^sup>2))" using \0 < \\ by (intro continuous_intros) (auto simp: sin_zero_pi_iff) show "(\x. inverse (2 * (sin (x * inverse 2))\<^sup>2)) \ borel_measurable (lebesgue_on {\..pi})" using \0 < \\ by (intro cont continuous_imp_measurable_on_sets_lebesgue) auto show "bounded ((\x. inverse (2 * (sin (x * inverse 2))\<^sup>2)) ` {\..pi})" using cont by (simp add: compact_continuous_image compact_imp_bounded) show "h absolutely_integrable_on {\..pi}" using \0 < \\ \\ < pi\ by (auto intro: absolutely_integrable_on_subinterval [OF h_aint]) qed auto then have *: "integrable (lebesgue_on {\..pi}) (\x. \h x\ / (2 * (sin (x/2))\<^sup>2))" by (simp add: absolutely_integrable_measurable o_def) define \ where "\ \ \x. (if n = 0 then 0 else if x = 0 then n/2 else (sin (real n / 2 * x))\<^sup>2 / (2 * real n * (sin (x/2))\<^sup>2)) * h x" have "\LINT x|lebesgue_on {\..pi}. \ x\ \ (LINT x|lebesgue_on {\..pi}. \h x\ / (2 * (sin (x/2))\<^sup>2) / n)" proof (rule integral_abs_bound_integral) show **: "integrable (lebesgue_on {\..pi}) (\x. \h x\ / (2 * (sin (x/2))\<^sup>2) / n)" using Bochner_Integration.integrable_mult_left [OF *, of "1/n"] by (simp add: field_simps) show \: "\\ x\ \ \h x\ / (2 * (sin (x/2))\<^sup>2) / real n" if "x \ space (lebesgue_on {\..pi})" for x using that \0 < \\ apply (simp add: \_def divide_simps mult_less_0_iff abs_mult) apply (auto simp: square_le_1 mult_left_le_one_le) done show "integrable (lebesgue_on {\..pi}) \" proof (rule measurable_bounded_by_integrable_imp_lebesgue_integrable [OF _ **]) let ?g = "\x. \h x\ / (2 * sin(x/2) ^ 2) / n" have ***: "integrable (lebesgue_on {\..pi}) (\x. (sin (n/2 * x))\<^sup>2 * (inverse (2 * (sin (x/2))\<^sup>2) * h x))" proof (rule absolutely_integrable_imp_integrable [OF absolutely_integrable_bounded_measurable_product_real]) show "(\x. (sin (real n / 2 * x))\<^sup>2) \ borel_measurable (lebesgue_on {\..pi})" by (intro continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto show "bounded ((\x. (sin (real n / 2 * x))\<^sup>2) ` {\..pi})" by (force simp: square_le_1 intro: boundedI [where B=1]) show "(\x. inverse (2 * (sin (x/2))\<^sup>2) * h x) absolutely_integrable_on {\..pi}" using hint by (simp add: divide_simps) qed auto show "\ \ borel_measurable (lebesgue_on {\..pi})" apply (rule borel_measurable_integrable) apply (rule integrable_cong [where f = "\x. sin(n / 2 * x) ^ 2 / (2 * n * sin(x/2) ^ 2) * h x", OF refl, THEN iffD1]) using \0 < \\ ** apply (force simp: \_def divide_simps algebra_simps mult_less_0_iff abs_mult) using Bochner_Integration.integrable_mult_left [OF ***, of "1/n"] by (simp add: field_simps) show "norm (\ x) \ ?g x" if "x \ {\..pi}" for x using that \ by (simp add: \_def) qed auto qed then show "norm (LINT x|lebesgue_on {\..pi}. Fejer_kernel n x * h x) \ \ n" by (simp add: Fejer_kernel \_def \_def flip: Bochner_Integration.integral_divide [OF *]) qed show "\ \ 0" unfolding \_def divide_inverse by (simp add: tendsto_mult_right_zero lim_inverse_n) qed then obtain N where N: "\n. n \ N \ \LINT x|lebesgue_on {\..pi}. Fejer_kernel n x * h x\ < e/2" unfolding lim_sequentially by (metis half_gt_zero_iff norm_conv_dist real_norm_def \e > 0\) show "\N. \n\N. \(LINT u|lebesgue_on {0..pi}. Fejer_kernel n u * h u)\ < e" proof (intro exI allI impI) fix n :: nat assume n: "n \ max 1 N" with N have 1: "\LINT x|lebesgue_on {\..pi}. Fejer_kernel n x * h x\ < e/2" by simp have "integral\<^sup>L (lebesgue_on {0..\}) (Fejer_kernel n) \ integral\<^sup>L (lebesgue_on {0..pi}) (Fejer_kernel n)" using \\ < pi\ has_bochner_integral_iff has_integral_Fejer_kernel_half by (force intro!: integral_mono_lebesgue_on_AE) also have "\ \ pi" using has_integral_Fejer_kernel_half by (simp add: has_bochner_integral_iff) finally have int_le_pi: "integral\<^sup>L (lebesgue_on {0..\}) (Fejer_kernel n) \ pi" . have 2: "\LINT x|lebesgue_on {0..\}. Fejer_kernel n x * h x\ \ (LINT x|lebesgue_on {0..\}. Fejer_kernel n x * e/2/pi)" proof - have eq: "integral\<^sup>L (lebesgue_on {0..\}) (\x. Fejer_kernel n x * h x) = integral\<^sup>L (lebesgue_on {0..\}) (\x. Fejer_kernel n x * (if x = 0 then 0 else h x))" proof (rule integral_cong_AE) have \: "{u. u \ 0 \ P u} \ {0..\} = {0} \ Collect P \ {0..\}" "{u. u \ 0 \ P u} \ {0..\} = (Collect P \ {0..\}) - {0}" for P using \0 < \\ by auto have *: "- {0} \ A \ {0..\} = A \ {0..\} - {0}" for A by auto show "(\x. Fejer_kernel n x * h x) \ borel_measurable (lebesgue_on {0..\})" using \\ < pi\ by (intro absolutely_integrable_imp_borel_measurable h_aint absolutely_integrable_on_subinterval [OF absolutely_integrable_mult_Fejer_kernel], auto) then show "(\x. Fejer_kernel n x * (if x = 0 then 0 else h x)) \ borel_measurable (lebesgue_on {0..\})" apply (simp add: in_borel_measurable Ball_def vimage_def Collect_conj_eq Collect_imp_eq * flip: Collect_neg_eq) apply (elim all_forward imp_forward asm_rl) using \0 < \\ - apply (auto simp: \ insert_sets_lebesgue_on cong: conj_cong) + apply (auto simp: \ sets.insert_in_sets sets_restrict_space_iff cong: conj_cong) done have 0: "{0} \ null_sets (lebesgue_on {0..\})" using \0 < \\ by (simp add: null_sets_restrict_space) then show "AE x in lebesgue_on {0..\}. Fejer_kernel n x * h x = Fejer_kernel n x * (if x = 0 then 0 else h x)" by (auto intro: AE_I' [OF 0]) qed show ?thesis unfolding eq proof (rule integral_abs_bound_integral) have "(\x. if x = 0 then 0 else h x) absolutely_integrable_on {- pi..pi}" proof (rule absolutely_integrable_spike [OF h_aint]) show "negligible {0}" by auto qed auto with \0 < \\ \\ < pi\ show "integrable (lebesgue_on {0..\}) (\x. Fejer_kernel n x * (if x = 0 then 0 else h x))" by (intro absolutely_integrable_imp_integrable h_aint absolutely_integrable_on_subinterval [OF absolutely_integrable_mult_Fejer_kernel]) auto show "integrable (lebesgue_on {0..\}) (\x. Fejer_kernel n x * e / 2 / pi)" by (simp add: absolutely_integrable_imp_integrable \\ < pi\ absolutely_integrable_mult_Fejer_kernel_reflected_part3 less_eq_real_def) show "\Fejer_kernel n x * (if x = 0 then 0 else h x)\ \ Fejer_kernel n x * e / 2 / pi" if "x \ space (lebesgue_on {0..\})" for x using that \ [of x] \e > 0\ by (auto simp: abs_mult eq simp flip: times_divide_eq_right intro: mult_left_mono) qed qed have 3: "\ \ e/2" using int_le_pi \0 < e\ by (simp add: divide_simps mult.commute [of e]) + have "integrable (lebesgue_on {0..pi}) (\x. Fejer_kernel n x * h x)" unfolding h_def by (simp add: absolutely_integrable_imp_integrable absolutely_integrable_mult_Fejer_kernel_reflected_part5 assms) then have "LINT x|lebesgue_on {0..pi}. Fejer_kernel n x * h x - = (LINT x|lebesgue_on {0..\}. Fejer_kernel n x * h x) + (LINT x|lebesgue_on {\..pi}. Fejer_kernel n x * h x)" + = (LINT x|lebesgue_on {0..\}. Fejer_kernel n x * h x) + (LINT x|lebesgue_on {\..pi}. Fejer_kernel n x * h x)" by (rule integral_combine) (use \0 < \\ \\ < pi\ in auto) then show "\LINT u|lebesgue_on {0..pi}. Fejer_kernel n u * h u\ < e" using 1 2 3 by linarith qed qed qed then show ?thesis unfolding h_def by (simp add: Fourier_sum_limit_Fejer_kernel_half assms add_divide_distrib) qed corollary Fourier_Fejer_Cesaro_summable_simple: -assumes f: "continuous_on UNIV f" + assumes f: "continuous_on UNIV f" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\mk\2*m. Fourier_coefficient f k * trigonometric_set k x) / n) \ f x" proof - have "(\n. (\mk\2*m. Fourier_coefficient f k * trigonometric_set k x) / n) \ (f x + f x) / 2" proof (rule Fourier_Fejer_Cesaro_summable) show "f absolutely_integrable_on {- pi..pi}" using absolutely_integrable_continuous_real continuous_on_subset f by blast show "(f \ f x) (at x within {..x})" "(f \ f x) (at x within {x..})" using Lim_at_imp_Lim_at_within continuous_on_def f by blast+ qed (auto simp: periodic Lim_at_imp_Lim_at_within continuous_on_def f) then show ?thesis by simp qed end diff --git a/thys/Fourier/Fourier_Aux2.thy b/thys/Fourier/Fourier_Aux2.thy --- a/thys/Fourier/Fourier_Aux2.thy +++ b/thys/Fourier/Fourier_Aux2.thy @@ -1,139 +1,42 @@ section\Lemmas possibly destined for future Isabelle releases\ theory Fourier_Aux2 - imports Ergodic_Theory.SG_Library_Complement + imports "HOL-Analysis.Analysis" begin - -lemma has_integral_sin_nx: "((\x. sin(real_of_int n * x)) has_integral 0) {-pi..pi}" -proof (cases "n = 0") - case False - have "((\x. sin (n * x)) has_integral (- cos (n * pi)/n - - cos (n * - pi)/n)) {-pi..pi}" - proof (rule fundamental_theorem_of_calculus) - show "((\x. - cos (n * x) / n) has_vector_derivative sin (n * a)) (at a within {-pi..pi})" - if "a \ {-pi..pi}" for a :: real - using that False - apply (simp only: has_vector_derivative_def) - apply (rule derivative_eq_intros | force)+ - done - qed auto - then show ?thesis - by simp -qed auto - -lemma integral_sin_nx: - "integral {-pi..pi} (\x. sin(x * real_of_int n)) = 0" - using has_integral_sin_nx [of n] by (force simp: mult.commute) - + lemma integral_sin_Z [simp]: assumes "n \ \" shows "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin(x * n)) = 0" proof (subst lebesgue_integral_eq_integral) show "integrable (lebesgue_on {-pi..pi}) (\x. sin (x * n))" by (intro continuous_imp_integrable_real continuous_intros) show "integral {-pi..pi} (\x. sin (x * n)) = 0" using assms Ints_cases integral_sin_nx by blast qed auto lemma integral_sin_Z' [simp]: assumes "n \ \" shows "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin(n * x)) = 0" by (metis assms integral_sin_Z mult_commute_abs) - -lemma has_integral_cos_nx: - "((\x. cos(real_of_int n * x)) has_integral (if n = 0 then 2 * pi else 0)) {-pi..pi}" -proof (cases "n = 0") - case True - then show ?thesis - using has_integral_const_real [of "1::real" "-pi" pi] by auto -next - case False - have "((\x. cos (n * x)) has_integral (sin (n * pi)/n - sin (n * - pi)/n)) {-pi..pi}" - proof (rule fundamental_theorem_of_calculus) - show "((\x. sin (n * x) / n) has_vector_derivative cos (n * x)) (at x within {-pi..pi})" - if "x \ {-pi..pi}" - for x :: real - using that False - apply (simp only: has_vector_derivative_def) - apply (rule derivative_eq_intros | force)+ - done - qed auto - with False show ?thesis - by (simp add: mult.commute) -qed - -lemma integral_cos_nx: - "integral {-pi..pi} (\x. cos(x * real_of_int n)) = (if n = 0 then 2 * pi else 0)" - using has_integral_cos_nx [of n] by (force simp: mult.commute) - + lemma integral_cos_Z [simp]: assumes "n \ \" shows "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(x * n)) = (if n = 0 then 2 * pi else 0)" proof (subst lebesgue_integral_eq_integral) show "integrable (lebesgue_on {-pi..pi}) (\x. cos (x * n))" by (intro continuous_imp_integrable_real continuous_intros) show "integral {-pi..pi} (\x. cos (x * n)) = (if n = 0 then 2 * pi else 0)" by (metis Ints_cases assms integral_cos_nx of_int_0_eq_iff) qed auto lemma integral_cos_Z' [simp]: assumes "n \ \" shows "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(n * x)) = (if n = 0 then 2 * pi else 0)" by (metis assms integral_cos_Z mult_commute_abs) -lemma real_tendsto_zero_mult_right_iff [simp]: - fixes c::real assumes "c \ 0" shows "(\n. a n * c)\ 0 \ a \ 0" - by (metis assms mult_zero_left real_tendsto_mult_right_iff) - -lemma real_tendsto_zero_divide_iff [simp]: - fixes c::real assumes "c \ 0" shows "(\n. a n / c)\ 0 \ a \ 0" - using real_tendsto_zero_mult_right_iff [of "1/c" a] assms by (simp add: field_simps) - -lemma insert_sets_lebesgue_on: - assumes "A \ sets (lebesgue_on S)" "a \ S" "S \ sets lebesgue" - shows "insert a A \ sets (lebesgue_on S)" - by (metis assms borel_singleton insert_subset lborelD sets.Int_space_eq2 sets.empty_sets sets.insert_in_sets sets_completionI_sets sets_restrict_space_iff) - lemma odd_even_cases [case_names 0 odd even]: assumes "P 0" and odd: "\n. P(Suc (2 * n))" and even: "\n. P(2 * n + 2)" shows "P n" by (metis nat_induct2 One_nat_def Suc_1 add_Suc_right assms(1) dvdE even odd oddE) -lemma measure_lebesgue_on_ivl [simp]: "\{a..b} \ S; S \ sets lebesgue\ \ measure (lebesgue_on S) {a..b} = content {a..b::real}" - by (simp add: measure_restrict_space) - - -lemma has_bochner_integral_null [intro]: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "N \ null_sets lebesgue" - shows "has_bochner_integral (lebesgue_on N) f 0" - unfolding has_bochner_integral_iff -proof - show "integrable (lebesgue_on N) f" - proof (subst integrable_restrict_space) - show "N \ space lebesgue \ sets lebesgue" - using assms by force - show "integrable lebesgue (\x. indicat_real N x *\<^sub>R f x)" - proof (rule integrable_cong_AE_imp) - show "integrable lebesgue (\x. 0)" - by simp - show *: "AE x in lebesgue. 0 = indicat_real N x *\<^sub>R f x" - using assms - by (simp add: indicator_def completion.null_sets_iff_AE eventually_mono) - show "(\x. indicat_real N x *\<^sub>R f x) \ borel_measurable lebesgue" - by (auto intro: borel_measurable_AE [OF _ *]) - qed - qed - show "integral\<^sup>L (lebesgue_on N) f = 0" - proof (rule integral_eq_zero_AE) - show "AE x in lebesgue_on N. f x = 0" - by (rule AE_I' [where N=N]) (auto simp: assms null_setsD2 null_sets_restrict_space) - qed -qed - -lemma has_bochner_integral_null_eq[simp]: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "N \ null_sets lebesgue" - shows "has_bochner_integral (lebesgue_on N) f i \ i = 0" - using assms has_bochner_integral_eq by blast - end diff --git a/thys/Landau_Symbols/Landau_Real_Products.thy b/thys/Landau_Symbols/Landau_Real_Products.thy --- a/thys/Landau_Symbols/Landau_Real_Products.thy +++ b/thys/Landau_Symbols/Landau_Real_Products.thy @@ -1,1483 +1,1483 @@ (* File: Landau_Real_Products.thy Author: Manuel Eberl Mathematical background and reification for a decision procedure for Landau symbols of products of powers of real functions (currently the identity and the natural logarithm) TODO: more functions (exp?), more preprocessing (log) *) section \Decision procedure for real functions\ theory Landau_Real_Products imports Main "HOL-Library.Function_Algebras" "HOL-Library.Set_Algebras" "HOL-Library.Landau_Symbols" Group_Sort begin subsection \Eventual non-negativity/non-zeroness\ text \ For certain transformations of Landau symbols, it is required that the functions involved are eventually non-negative of non-zero. In the following, we set up a system to guide the simplifier to discharge these requirements during simplification at least in obvious cases. \ definition "eventually_nonzero F f \ eventually (\x. (f x :: _ :: real_normed_field) \ 0) F" definition "eventually_nonneg F f \ eventually (\x. (f x :: _ :: linordered_field) \ 0) F" named_theorems eventually_nonzero_simps lemmas [eventually_nonzero_simps] = eventually_nonzero_def [symmetric] eventually_nonneg_def [symmetric] lemma eventually_nonzeroD: "eventually_nonzero F f \ eventually (\x. f x \ 0) F" by (simp add: eventually_nonzero_def) lemma eventually_nonzero_const [eventually_nonzero_simps]: "eventually_nonzero F (\_::_::linorder. c) \ F = bot \ c \ 0" unfolding eventually_nonzero_def by (auto simp add: eventually_False) lemma eventually_nonzero_inverse [eventually_nonzero_simps]: "eventually_nonzero F (\x. inverse (f x)) \ eventually_nonzero F f" unfolding eventually_nonzero_def by simp lemma eventually_nonzero_mult [eventually_nonzero_simps]: "eventually_nonzero F (\x. f x * g x) \ eventually_nonzero F f \ eventually_nonzero F g" unfolding eventually_nonzero_def by (simp_all add: eventually_conj_iff[symmetric]) lemma eventually_nonzero_pow [eventually_nonzero_simps]: "eventually_nonzero F (\x::_::linorder. f x ^ n) \ n = 0 \ eventually_nonzero F f" by (induction n) (auto simp: eventually_nonzero_simps) lemma eventually_nonzero_divide [eventually_nonzero_simps]: "eventually_nonzero F (\x. f x / g x) \ eventually_nonzero F f \ eventually_nonzero F g" unfolding eventually_nonzero_def by (simp_all add: eventually_conj_iff[symmetric]) lemma eventually_nonzero_ident_at_top_linorder [eventually_nonzero_simps]: "eventually_nonzero at_top (\x::'a::{real_normed_field,linordered_field}. x)" unfolding eventually_nonzero_def by simp lemma eventually_nonzero_ident_nhds [eventually_nonzero_simps]: "eventually_nonzero (nhds a) (\x. x) \ a \ 0" using eventually_nhds_in_open[of "-{0}" a] by (auto elim!: eventually_mono simp: eventually_nonzero_def open_Compl dest: eventually_nhds_x_imp_x) lemma eventually_nonzero_ident_at_within [eventually_nonzero_simps]: "eventually_nonzero (at a within A) (\x. x)" using eventually_nonzero_ident_nhds[of a] by (cases "a = 0") (auto simp: eventually_nonzero_def eventually_at_filter elim!: eventually_mono) lemma eventually_nonzero_ln_at_top [eventually_nonzero_simps]: "eventually_nonzero at_top (\x::real. ln x)" unfolding eventually_nonzero_def by (auto intro!: eventually_mono[OF eventually_gt_at_top[of 1]]) lemma eventually_nonzero_ln_const_at_top [eventually_nonzero_simps]: "b > 0 \ eventually_nonzero at_top (\x. ln (b * x :: real))" unfolding eventually_nonzero_def apply (rule eventually_mono [OF eventually_gt_at_top[of "max 1 (inverse b)"]]) by (metis exp_ln exp_minus exp_minus_inverse less_numeral_extra(3) ln_gt_zero max_less_iff_conj mult.commute mult_strict_right_mono) lemma eventually_nonzero_ln_const'_at_top [eventually_nonzero_simps]: "b > 0 \ eventually_nonzero at_top (\x. ln (x * b :: real))" using eventually_nonzero_ln_const_at_top[of b] by (simp add: mult.commute) lemma eventually_nonzero_powr_at_top [eventually_nonzero_simps]: "eventually_nonzero at_top (\x::real. f x powr p) \ eventually_nonzero at_top f" unfolding eventually_nonzero_def by simp lemma eventually_nonneg_const [eventually_nonzero_simps]: "eventually_nonneg F (\_. c) \ F = bot \ c \ 0" unfolding eventually_nonneg_def by (auto simp: eventually_False) lemma eventually_nonneg_inverse [eventually_nonzero_simps]: "eventually_nonneg F (\x. inverse (f x)) \ eventually_nonneg F f" unfolding eventually_nonneg_def by (intro eventually_subst) (auto) lemma eventually_nonneg_add [eventually_nonzero_simps]: assumes "eventually_nonneg F f" "eventually_nonneg F g" shows "eventually_nonneg F (\x. f x + g x)" using assms unfolding eventually_nonneg_def by eventually_elim simp lemma eventually_nonneg_mult [eventually_nonzero_simps]: assumes "eventually_nonneg F f" "eventually_nonneg F g" shows "eventually_nonneg F (\x. f x * g x)" using assms unfolding eventually_nonneg_def by eventually_elim simp lemma eventually_nonneg_mult' [eventually_nonzero_simps]: assumes "eventually_nonneg F (\x. -f x)" "eventually_nonneg F (\x. - g x)" shows "eventually_nonneg F (\x. f x * g x)" using assms unfolding eventually_nonneg_def by eventually_elim (auto intro: mult_nonpos_nonpos) lemma eventually_nonneg_divide [eventually_nonzero_simps]: assumes "eventually_nonneg F f" "eventually_nonneg F g" shows "eventually_nonneg F (\x. f x / g x)" using assms unfolding eventually_nonneg_def by eventually_elim simp lemma eventually_nonneg_divide' [eventually_nonzero_simps]: assumes "eventually_nonneg F (\x. -f x)" "eventually_nonneg F (\x. - g x)" shows "eventually_nonneg F (\x. f x / g x)" using assms unfolding eventually_nonneg_def by eventually_elim (auto intro: divide_nonpos_nonpos) lemma eventually_nonneg_ident_at_top [eventually_nonzero_simps]: "eventually_nonneg at_top (\x. x)" unfolding eventually_nonneg_def by (rule eventually_ge_at_top) lemma eventually_nonneg_ident_nhds [eventually_nonzero_simps]: fixes a :: "'a :: {linorder_topology, linordered_field}" shows "a > 0 \ eventually_nonneg (nhds a) (\x. x)" unfolding eventually_nonneg_def using eventually_nhds_in_open[of "{0<..}" a] by (auto simp: eventually_nonneg_def dest: eventually_nhds_x_imp_x elim!: eventually_mono) lemma eventually_nonneg_ident_at_within [eventually_nonzero_simps]: fixes a :: "'a :: {linorder_topology, linordered_field}" shows "a > 0 \ eventually_nonneg (at a within A) (\x. x)" using eventually_nonneg_ident_nhds[of a] by (auto simp: eventually_nonneg_def eventually_at_filter elim: eventually_mono) lemma eventually_nonneg_pow [eventually_nonzero_simps]: "eventually_nonneg F f \ eventually_nonneg F (\x. f x ^ n)" by (induction n) (auto simp: eventually_nonzero_simps) lemma eventually_nonneg_powr [eventually_nonzero_simps]: "eventually_nonneg F (\x. f x powr y :: real)" by (simp add: eventually_nonneg_def) lemma eventually_nonneg_ln_at_top [eventually_nonzero_simps]: "eventually_nonneg at_top (\x. ln x :: real)" by (auto intro!: eventually_mono[OF eventually_gt_at_top[of "1::real"]] simp: eventually_nonneg_def) lemma eventually_nonneg_ln_const [eventually_nonzero_simps]: "b > 0 \ eventually_nonneg at_top (\x. ln (b*x) :: real)" unfolding eventually_nonneg_def using eventually_ge_at_top[of "inverse b"] by eventually_elim (simp_all add: field_simps) lemma eventually_nonneg_ln_const' [eventually_nonzero_simps]: "b > 0 \ eventually_nonneg at_top (\x. ln (x*b) :: real)" using eventually_nonneg_ln_const[of b] by (simp add: mult.commute) lemma eventually_nonzero_bigtheta': "f \ \[F](g) \ eventually_nonzero F f \ eventually_nonzero F g" unfolding eventually_nonzero_def by (rule eventually_nonzero_bigtheta) lemma eventually_nonneg_at_top: assumes "filterlim f at_top F" shows "eventually_nonneg F f" proof - from assms have "eventually (\x. f x \ 0) F" by (simp add: filterlim_at_top) thus ?thesis unfolding eventually_nonneg_def by eventually_elim simp qed lemma eventually_nonzero_at_top: assumes "filterlim (f :: 'a \ 'b :: {linordered_field, real_normed_field}) at_top F" shows "eventually_nonzero F f" proof - from assms have "eventually (\x. f x \ 1) F" by (simp add: filterlim_at_top) thus ?thesis unfolding eventually_nonzero_def by eventually_elim auto qed lemma eventually_nonneg_at_top_ASSUMPTION [eventually_nonzero_simps]: "ASSUMPTION (filterlim f at_top F) \ eventually_nonneg F f" by (simp add: ASSUMPTION_def eventually_nonneg_at_top) lemma eventually_nonzero_at_top_ASSUMPTION [eventually_nonzero_simps]: "ASSUMPTION (filterlim f (at_top :: 'a :: {linordered_field, real_normed_field} filter) F) \ eventually_nonzero F f" using eventually_nonzero_at_top[of f F] by (simp add: ASSUMPTION_def) lemma filterlim_at_top_iff_smallomega: fixes f :: "_ \ real" shows "filterlim f at_top F \ f \ \[F](\_. 1) \ eventually_nonneg F f" unfolding eventually_nonneg_def proof safe assume A: "filterlim f at_top F" thus B: "eventually (\x. f x \ 0) F" by (simp add: eventually_nonzero_simps) { fix c from A have "filterlim (\x. norm (f x)) at_top F" by (intro filterlim_at_infinity_imp_norm_at_top filterlim_at_top_imp_at_infinity) hence "eventually (\x. norm (f x) \ c) F" by (auto simp: filterlim_at_top) } thus "f \ \[F](\_. 1)" by (rule landau_omega.smallI) next assume A: "f \ \[F](\_. 1)" and B: "eventually (\x. f x \ 0) F" { fix c :: real assume "c > 0" from landau_omega.smallD[OF A this] B have "eventually (\x. f x \ c) F" by eventually_elim simp } thus "filterlim f at_top F" by (subst filterlim_at_top_gt[of _ _ 0]) simp_all qed lemma smallomega_1_iff: "eventually_nonneg F f \ f \ \[F](\_. 1 :: real) \ filterlim f at_top F" by (simp add: filterlim_at_top_iff_smallomega) lemma smallo_1_iff: "eventually_nonneg F f \ (\_. 1 :: real) \ o[F](f) \ filterlim f at_top F" by (simp add: filterlim_at_top_iff_smallomega smallomega_iff_smallo) lemma eventually_nonneg_add1 [eventually_nonzero_simps]: assumes "eventually_nonneg F f" "g \ o[F](f)" shows "eventually_nonneg F (\x. f x + g x :: real)" using landau_o.smallD[OF assms(2) zero_less_one] assms(1) unfolding eventually_nonneg_def by eventually_elim simp_all lemma eventually_nonneg_add2 [eventually_nonzero_simps]: assumes "eventually_nonneg F g" "f \ o[F](g)" shows "eventually_nonneg F (\x. f x + g x :: real)" using landau_o.smallD[OF assms(2) zero_less_one] assms(1) unfolding eventually_nonneg_def by eventually_elim simp_all lemma eventually_nonneg_diff1 [eventually_nonzero_simps]: assumes "eventually_nonneg F f" "g \ o[F](f)" shows "eventually_nonneg F (\x. f x - g x :: real)" using landau_o.smallD[OF assms(2) zero_less_one] assms(1) unfolding eventually_nonneg_def by eventually_elim simp_all lemma eventually_nonneg_diff2 [eventually_nonzero_simps]: assumes "eventually_nonneg F (\x. - g x)" "f \ o[F](g)" shows "eventually_nonneg F (\x. f x - g x :: real)" using landau_o.smallD[OF assms(2) zero_less_one] assms(1) unfolding eventually_nonneg_def by eventually_elim simp_all subsection \Rewriting Landau symbols\ lemma bigtheta_mult_eq: "\[F](\x. f x * g x) = \[F](f) * \[F](g)" proof (intro equalityI subsetI) fix h assume "h \ \[F](f) * \[F](g)" thus "h \ \[F](\x. f x * g x)" by (elim set_times_elim, hypsubst, unfold func_times) (erule (1) landau_theta.mult) next fix h assume "h \ \[F](\x. f x * g x)" then guess c1 c2 :: real unfolding bigtheta_def by (elim landau_o.bigE landau_omega.bigE IntE) note c = this define h1 h2 where "h1 x = (if g x = 0 then if f x = 0 then if h x = 0 then h x else 1 else f x else h x / g x)" and "h2 x = (if g x = 0 then if f x = 0 then h x else h x / f x else g x)" for x have "h = h1 * h2" by (intro ext) (auto simp: h1_def h2_def field_simps) moreover have "h1 \ \[F](f)" proof (rule bigthetaI') from c(3) show "min c2 1 > 0" by simp from c(1) show "max c1 1 > 0" by simp from c(2,4) show "eventually (\x. min c2 1 * (norm (f x)) \ norm (h1 x) \ norm (h1 x) \ max c1 1 * (norm (f x))) F" apply eventually_elim proof (rule conjI) fix x assume A: "(norm (h x)) \ c1 * norm (f x * g x)" and B: "(norm (h x)) \ c2 * norm (f x * g x)" have m: "min c2 1 * (norm (f x)) \ 1 * (norm (f x))" by (rule mult_right_mono) simp_all have "min c2 1 * norm (f x * g x) \ c2 * norm (f x * g x)" by (intro mult_right_mono) simp_all also note B finally show "norm (h1 x) \ min c2 1 * (norm (f x))" using m A by (cases "g x = 0") (simp_all add: h1_def norm_mult norm_divide field_simps)+ have m: "1 * (norm (f x)) \ max c1 1 * (norm (f x))" by (rule mult_right_mono) simp_all note A also have "c1 * norm (f x * g x) \ max c1 1 * norm (f x * g x)" by (intro mult_right_mono) simp_all finally show "norm (h1 x) \ max c1 1 * (norm (f x))" using m A by (cases "g x = 0") (simp_all add: h1_def norm_mult norm_divide field_simps)+ qed qed moreover have "h2 \ \[F](g)" proof (rule bigthetaI') from c(3) show "min c2 1 > 0" by simp from c(1) show "max c1 1 > 0" by simp from c(2,4) show "eventually (\x. min c2 1 * (norm (g x)) \ norm (h2 x) \ norm (h2 x) \ max c1 1 * (norm (g x))) F" apply eventually_elim proof (rule conjI) fix x assume A: "(norm (h x)) \ c1 * norm (f x * g x)" and B: "(norm (h x)) \ c2 * norm (f x * g x)" have m: "min c2 1 * (norm (f x)) \ 1 * (norm (f x))" by (rule mult_right_mono) simp_all have "min c2 1 * norm (f x * g x) \ c2 * norm (f x * g x)" by (intro mult_right_mono) simp_all also note B finally show "norm (h2 x) \ min c2 1 * (norm (g x))" using m A B by (cases "g x = 0") (auto simp: h2_def abs_mult field_simps)+ have m: "1 * (norm (g x)) \ max c1 1 * (norm (g x))" by (rule mult_right_mono) simp_all note A also have "c1 * norm (f x * g x) \ max c1 1 * norm (f x * g x)" by (intro mult_right_mono) simp_all finally show "norm (h2 x) \ max c1 1 * (norm (g x))" using m A by (cases "g x = 0") (simp_all add: h2_def abs_mult field_simps)+ qed qed ultimately show "h \ \[F](f) * \[F](g)" by blast qed text \ Since the simplifier does not currently rewriting with relations other than equality, but we want to rewrite terms like @{term "\(\x. log 2 x * x)"} to @{term "\(\x. ln x * x)"}, we need to bring the term into something that contains @{term "\(\x. log 2 x)"} and @{term "\(\x. x)"}, which can then be rewritten individually. For this, we introduce the following constants and rewrite rules. The rules are mainly used by the simprocs, but may be useful for manual reasoning occasionally. \ definition "set_mult A B = {\x. f x * g x |f g. f \ A \ g \ B}" definition "set_inverse A = {\x. inverse (f x) |f. f \ A}" definition "set_divide A B = {\x. f x / g x |f g. f \ A \ g \ B}" definition "set_pow A n = {\x. f x ^ n |f. f \ A}" definition "set_powr A y = {\x. f x powr y |f. f \ A}" lemma bigtheta_mult_eq_set_mult: shows "\[F](\x. f x * g x) = set_mult (\[F](f)) (\[F](g))" unfolding bigtheta_mult_eq set_mult_def set_times_def func_times by blast lemma bigtheta_inverse_eq_set_inverse: shows "\[F](\x. inverse (f x)) = set_inverse (\[F](f))" proof (intro equalityI subsetI) fix g :: "'a \ 'b" assume "g \ \[F](\x. inverse (f x))" hence "(\x. inverse (g x)) \ \[F](\x. inverse (inverse (f x)))" by (subst bigtheta_inverse) also have "(\x. inverse (inverse (f x))) = f" by (rule ext) simp finally show "g \ set_inverse (\[F](f))" unfolding set_inverse_def by force next fix g :: "'a \ 'b" assume "g \ set_inverse (\[F](f))" then obtain g' where "g = (\x. inverse (g' x))" "g' \ \[F](f)" unfolding set_inverse_def by blast hence "(\x. inverse (g' x)) \ \[F](\x. inverse (f x))" by (subst bigtheta_inverse) also from \g = (\x. inverse (g' x))\ have "(\x. inverse (g' x)) = g" by (intro ext) simp finally show "g \ \[F](\x. inverse (f x))" . qed lemma set_divide_inverse: "set_divide (A :: (_ \ (_ :: division_ring)) set) B = set_mult A (set_inverse B)" proof (intro equalityI subsetI) fix f assume "f \ set_divide A B" then obtain g h where "f = (\x. g x / h x)" "g \ A" "h \ B" unfolding set_divide_def by blast hence "f = g * (\x. inverse (h x))" "(\x. inverse (h x)) \ set_inverse B" unfolding set_inverse_def by (auto simp: divide_inverse) with \g \ A\ show "f \ set_mult A (set_inverse B)" unfolding set_mult_def by force next fix f assume "f \ set_mult A (set_inverse B)" then obtain g h where "f = g * (\x. inverse (h x))" "g \ A" "h \ B" unfolding set_times_def set_inverse_def set_mult_def by force hence "f = (\x. g x / h x)" by (intro ext) (simp add: divide_inverse) with \g \ A\ \h \ B\ show "f \ set_divide A B" unfolding set_divide_def by blast qed lemma bigtheta_divide_eq_set_divide: shows "\[F](\x. f x / g x) = set_divide (\[F](f)) (\[F](g))" by (simp only: set_divide_inverse divide_inverse bigtheta_mult_eq_set_mult bigtheta_inverse_eq_set_inverse) primrec bigtheta_pow where "bigtheta_pow F A 0 = \[F](\_. 1)" | "bigtheta_pow F A (Suc n) = set_mult A (bigtheta_pow F A n)" lemma bigtheta_pow_eq_set_pow: "\[F](\x. f x ^ n) = bigtheta_pow F (\[F](f)) n" by (induction n) (simp_all add: bigtheta_mult_eq_set_mult) definition bigtheta_powr where "bigtheta_powr F A y = (if y = 0 then {f. \g\A. eventually_nonneg F g \ f \ \[F](\x. g x powr y)} else {f. \g\A. eventually_nonneg F g \ (\x. (norm (f x)) = g x powr y)})" lemma bigtheta_powr_eq_set_powr: assumes "eventually_nonneg F f" shows "\[F](\x. f x powr (y::real)) = bigtheta_powr F (\[F](f)) y" proof (cases "y = 0") assume [simp]: "y = 0" show ?thesis proof (intro equalityI subsetI) fix h assume "h \ bigtheta_powr F \[F](f) y" then obtain g where g: "g \ \[F](f)" "eventually_nonneg F g" "h \ \[F](\x. g x powr 0)" unfolding bigtheta_powr_def by force note this(3) also have "(\x. g x powr 0) \ \[F](\x. \g x\ powr 0)" using assms unfolding eventually_nonneg_def by (intro bigthetaI_cong) (auto elim!: eventually_mono) also from g(1) have "(\x. \g x\ powr 0) \ \[F](\x. \f x\ powr 0)" by (rule bigtheta_powr) also from g(2) have "(\x. f x powr 0) \ \[F](\x. \f x\ powr 0)" unfolding eventually_nonneg_def by (intro bigthetaI_cong) (auto elim!: eventually_mono) finally show "h \ \[F](\x. f x powr y)" by simp next fix h assume "h \ \[F](\x. f x powr y)" with assms have "\g\\[F](f). eventually_nonneg F g \ h \ \[F](\x. g x powr 0)" by (intro bexI[of _ f] conjI) simp_all thus "h \ bigtheta_powr F \[F](f) y" unfolding bigtheta_powr_def by simp qed next assume y: "y \ 0" show ?thesis proof (intro equalityI subsetI) fix h assume h: "h \ \[F](\x. f x powr y)" let ?h' = "\x. \h x\ powr inverse y" from bigtheta_powr[OF h, of "inverse y"] y have "?h' \ \[F](\x. f x powr 1)" by (simp add: powr_powr) also have "(\x. f x powr 1) \ \[F](f)" using assms unfolding eventually_nonneg_def by (intro bigthetaI_cong) (auto elim!: eventually_mono) finally have "?h' \ \[F](f)" . with y have "\g\\[F](f). eventually_nonneg F g \ (\x. (norm (h x)) = g x powr y)" by (intro bexI[of _ ?h']) (simp_all add: powr_powr eventually_nonneg_def) thus "h \ bigtheta_powr F \[F](f) y" using y unfolding bigtheta_powr_def by simp next fix h assume "h \ bigtheta_powr F (\[F](f)) y" with y obtain g where A: "g \ \[F](f)" "\x. \h x\ = g x powr y" "eventually_nonneg F g" unfolding bigtheta_powr_def by force from this(3) have "(\x. g x powr y) \ \[F](\x. \g x\ powr y)" unfolding eventually_nonneg_def by (intro bigthetaI_cong) (auto elim!: eventually_mono) also from A(1) have "(\x. \g x\ powr y) \ \[F](\x. \f x\ powr y)" by (rule bigtheta_powr) also have "(\x. \f x\ powr y) \ \[F](\x. f x powr y)" using assms unfolding eventually_nonneg_def by (intro bigthetaI_cong) (auto elim!: eventually_mono) finally have "(\x. \h x\) \ \[F](\x. f x powr y)" by (subst A(2)) thus "(\x. h x) \ \[F](\x. f x powr y)" by simp qed qed lemmas bigtheta_factors_eq = bigtheta_mult_eq_set_mult bigtheta_inverse_eq_set_inverse bigtheta_divide_eq_set_divide bigtheta_pow_eq_set_pow bigtheta_powr_eq_set_powr lemmas landau_bigtheta_congs = landau_symbols[THEN landau_symbol.cong_bigtheta] lemma (in landau_symbol) meta_cong_bigtheta: "\[F](f) \ \[F](g) \ L F (f) \ L F (g)" using bigtheta_refl[of f] by (intro eq_reflection cong_bigtheta) blast lemmas landau_bigtheta_meta_congs = landau_symbols[THEN landau_symbol.meta_cong_bigtheta] subsection \Preliminary facts\ lemma real_powr_at_top: assumes "(p::real) > 0" shows "filterlim (\x. x powr p) at_top at_top" proof (subst filterlim_cong[OF refl refl]) show "LIM x at_top. exp (p * ln x) :> at_top" by (rule filterlim_compose[OF exp_at_top filterlim_tendsto_pos_mult_at_top[OF tendsto_const]]) (simp_all add: ln_at_top assms) show "eventually (\x. x powr p = exp (p * ln x)) at_top" using eventually_gt_at_top[of 0] by eventually_elim (simp add: powr_def) qed lemma tendsto_ln_over_powr: assumes "(a::real) > 0" shows "((\x. ln x / x powr a) \ 0) at_top" proof (rule lhospital_at_top_at_top) from assms show "LIM x at_top. x powr a :> at_top" by (rule real_powr_at_top) show "eventually (\x. a * x powr (a - 1) \ 0) at_top" using eventually_gt_at_top[of "0::real"] by eventually_elim (insert assms, simp) show "eventually (\x::real. (ln has_real_derivative (inverse x)) (at x)) at_top" using eventually_gt_at_top[of "0::real"] DERIV_ln by (elim eventually_mono) simp show "eventually (\x. ((\x. x powr a) has_real_derivative a * x powr (a - 1)) (at x)) at_top" using eventually_gt_at_top[of "0::real"] by eventually_elim (auto intro!: derivative_eq_intros) have "eventually (\x. inverse a * x powr -a = inverse x / (a*x powr (a-1))) at_top" using eventually_gt_at_top[of "0::real"] by (elim eventually_mono) (simp add: field_simps powr_diff powr_minus) moreover from assms have "((\x. inverse a * x powr -a) \ 0) at_top" by (intro tendsto_mult_right_zero tendsto_neg_powr filterlim_ident) simp_all ultimately show "((\x. inverse x / (a * x powr (a - 1))) \ 0) at_top" by (subst (asm) tendsto_cong) simp_all qed lemma tendsto_ln_powr_over_powr: assumes "(a::real) > 0" "b > 0" shows "((\x. ln x powr a / x powr b) \ 0) at_top" proof- have "eventually (\x. ln x powr a / x powr b = (ln x / x powr (b/a)) powr a) at_top" using assms eventually_gt_at_top[of "1::real"] by (elim eventually_mono) (simp add: powr_divide powr_powr) moreover have "eventually (\x. 0 < ln x / x powr (b / a)) at_top" using eventually_gt_at_top[of "1::real"] by (elim eventually_mono) simp with assms have "((\x. (ln x / x powr (b/a)) powr a) \ 0) at_top" by (intro tendsto_zero_powrI tendsto_ln_over_powr) (simp_all add: eventually_mono) ultimately show ?thesis by (subst tendsto_cong) simp_all qed lemma tendsto_ln_powr_over_powr': assumes "b > 0" shows "((\x::real. ln x powr a / x powr b) \ 0) at_top" proof (cases "a \ 0") assume a: "a \ 0" show ?thesis proof (rule tendsto_sandwich[of "\_::real. 0"]) have "eventually (\x. ln x powr a \ 1) at_top" unfolding eventually_at_top_linorder proof (intro allI exI impI) fix x :: real assume x: "x \ exp 1" have "0 < exp (1::real)" by simp also have "\ \ x" by fact finally have "ln x \ ln (exp 1)" using x by (subst ln_le_cancel_iff) auto hence "ln x powr a \ ln (exp 1) powr a" using a by (intro powr_mono2') simp_all thus "ln x powr a \ 1" by simp qed thus "eventually (\x. ln x powr a / x powr b \ x powr -b) at_top" by eventually_elim (insert a, simp add: field_simps powr_minus divide_right_mono) qed (auto intro!: filterlim_ident tendsto_neg_powr assms) qed (intro tendsto_ln_powr_over_powr, simp_all add: assms) lemma tendsto_ln_over_ln: assumes "(a::real) > 0" "c > 0" shows "((\x. ln (a*x) / ln (c*x)) \ 1) at_top" proof (rule lhospital_at_top_at_top) show "LIM x at_top. ln (c*x) :> at_top" by (intro filterlim_compose[OF ln_at_top] filterlim_tendsto_pos_mult_at_top[OF tendsto_const] filterlim_ident assms(2)) show "eventually (\x. ((\x. ln (a*x)) has_real_derivative (inverse x)) (at x)) at_top" using eventually_gt_at_top[of "inverse a"] assms by (auto elim!: eventually_mono intro!: derivative_eq_intros simp: field_simps) show "eventually (\x. ((\x. ln (c*x)) has_real_derivative (inverse x)) (at x)) at_top" using eventually_gt_at_top[of "inverse c"] assms by (auto elim!: eventually_mono intro!: derivative_eq_intros simp: field_simps) show "((\x::real. inverse x / inverse x) \ 1) at_top" by (subst tendsto_cong[of _ "\_. 1"]) simp_all qed simp_all lemma tendsto_ln_powr_over_ln_powr: assumes "(a::real) > 0" "c > 0" shows "((\x. ln (a*x) powr d / ln (c*x) powr d) \ 1) at_top" proof- have "eventually (\x. ln (a*x) powr d / ln (c*x) powr d = (ln (a*x) / ln (c*x)) powr d) at_top" using assms eventually_gt_at_top[of "max (inverse a) (inverse c)"] by (auto elim!: eventually_mono simp: powr_divide field_simps) moreover have "((\x. (ln (a*x) / ln (c*x)) powr d) \ 1) at_top" using assms by (intro tendsto_eq_rhs[OF tendsto_powr[OF tendsto_ln_over_ln tendsto_const]]) simp_all ultimately show ?thesis by (subst tendsto_cong) qed lemma tendsto_ln_powr_over_ln_powr': "c > 0 \ ((\x::real. ln x powr d / ln (c*x) powr d) \ 1) at_top" using tendsto_ln_powr_over_ln_powr[of 1 c d] by simp lemma tendsto_ln_powr_over_ln_powr'': "a > 0 \ ((\x::real. ln (a*x) powr d / ln x powr d) \ 1) at_top" using tendsto_ln_powr_over_ln_powr[of _ 1] by simp lemma bigtheta_const_ln_powr [simp]: "a > 0 \ (\x::real. ln (a*x) powr d) \ \(\x. ln x powr d)" by (intro bigthetaI_tendsto[of 1] tendsto_ln_powr_over_ln_powr'') simp lemma bigtheta_const_ln_pow [simp]: "a > 0 \ (\x::real. ln (a*x) ^ d) \ \(\x. ln x ^ d)" proof- assume a: "a > 0" have "\\<^sub>F x in at_top. ln (a * x) ^ d = ln (a * x) powr real d" using eventually_gt_at_top[of "1/a"] by eventually_elim (insert a, subst powr_realpow, auto simp: field_simps) hence "(\x::real. ln (a*x) ^ d) \ \(\x. ln (a*x) powr real d)" by (rule bigthetaI_cong) also from a have "(\x. ln (a*x) powr real d) \ \(\x. ln x powr real d)" by simp also have "\\<^sub>F x in at_top. ln x powr real d = ln x ^ d" using eventually_gt_at_top[of 1] by eventually_elim (subst powr_realpow, auto simp: field_simps) hence "(\x. ln x powr real d) \ \(\x. ln x ^ d)" by (rule bigthetaI_cong) finally show ?thesis . qed lemma bigtheta_const_ln [simp]: "a > 0 \ (\x::real. ln (a*x)) \ \(\x. ln x)" using tendsto_ln_over_ln[of a 1] by (intro bigthetaI_tendsto[of 1]) simp_all text \ If there are two functions @{term "f"} and @{term "g"} where any power of @{term "g"} is asymptotically smaller than @{term "f"}, propositions like @{term "(\x. f x ^ p1 * g x ^ q1) \ O(\x. f x ^ p2 * g x ^ q2)"} can be decided just by looking at the exponents: the proposition is true iff @{term "p1 < p2"} or @{term "p1 = p2 \ q1 \ q2"}. The functions @{term "\x. x"}, @{term "\x. ln x"}, @{term "\x. ln (ln x)"}, $\ldots$ form a chain in which every function dominates all succeeding functions in the above sense, allowing to decide propositions involving Landau symbols and functions that are products of powers of functions from this chain by reducing the proposition to a statement involving only logical connectives and comparisons on the exponents. We will now give the mathematical background for this and implement reification to bring functions from this class into a canonical form, allowing the decision procedure to be implemented in a simproc. \ subsection \Decision procedure\ definition "powr_closure f \ {\x. f x powr p :: real |p. True}" lemma powr_closureI [simp]: "(\x. f x powr p) \ powr_closure f" unfolding powr_closure_def by force lemma powr_closureE: assumes "g \ powr_closure f" obtains p where "g = (\x. f x powr p)" using assms unfolding powr_closure_def by force locale landau_function_family = fixes F :: "'a filter" and H :: "('a \ real) set" assumes F_nontrivial: "F \ bot" assumes pos: "h \ H \ eventually (\x. h x > 0) F" assumes linear: "h1 \ H \ h2 \ H \ h1 \ o[F](h2) \ h2 \ o[F](h1) \ h1 \ \[F](h2)" assumes mult: "h1 \ H \ h2 \ H \ (\x. h1 x * h2 x) \ H" assumes inverse: "h \ H \ (\x. inverse (h x)) \ H" begin lemma div: "h1 \ H \ h2 \ H \ (\x. h1 x / h2 x) \ H" by (subst divide_inverse) (intro mult inverse) lemma nonzero: "h \ H \ eventually (\x. h x \ 0) F" by (drule pos) (auto elim: eventually_mono) lemma landau_cases: assumes "h1 \ H" "h2 \ H" obtains "h1 \ o[F](h2)" | "h2 \ o[F](h1)" | "h1 \ \[F](h2)" using linear[OF assms] by blast lemma small_big_antisym: assumes "h1 \ H" "h2 \ H" "h1 \ o[F](h2)" "h2 \ O[F](h1)" shows False proof- from nonzero[OF assms(1)] nonzero[OF assms(2)] landau_o.small_big_asymmetric[OF assms(3,4)] have "eventually (\_::'a. False) F" by eventually_elim simp thus False by (simp add: eventually_False F_nontrivial) qed lemma small_antisym: assumes "h1 \ H" "h2 \ H" "h1 \ o[F](h2)" "h2 \ o[F](h1)" shows False using assms by (blast intro: small_big_antisym landau_o.small_imp_big) end locale landau_function_family_pair = G: landau_function_family F G + H: landau_function_family F H for F G H + fixes g assumes gs_dominate: "g1 \ G \ g2 \ G \ h1 \ H \ h2 \ H \ g1 \ o[F](g2) \ (\x. g1 x * h1 x) \ o[F](\x. g2 x * h2 x)" assumes g: "g \ G" assumes g_dominates: "h \ H \ h \ o[F](g)" begin sublocale GH: landau_function_family F "G * H" proof (unfold_locales; (elim set_times_elim; hypsubst)?) fix g h assume "g \ G" "h \ H" from G.pos[OF this(1)] H.pos[OF this(2)] show "eventually (\x. (g*h) x > 0) F" by eventually_elim simp next fix g h assume A: "g \ G" "h \ H" have "(\x. inverse ((g * h) x)) = (\x. inverse (g x)) * (\x. inverse (h x))" by (rule ext) simp also from A have "... \ G * H" by (intro G.inverse H.inverse set_times_intro) finally show "(\x. inverse ((g * h) x)) \ G * H" . next fix g1 g2 h1 h2 assume A: "g1 \ G" "g2 \ G" "h1 \ H" "h2 \ H" from gs_dominate[OF this] gs_dominate[OF this(2,1,4,3)] G.linear[OF this(1,2)] H.linear[OF this(3,4)] show "g1 * h1 \ o[F](g2 * h2) \ g2 * h2 \ o[F](g1 * h1) \ g1 * h1 \ \[F](g2 * h2)" by (elim disjE) (force simp: func_times bigomega_iff_bigo intro: landau_theta.mult landau_o.small.mult landau_o.small_big_mult landau_o.big_small_mult)+ have B: "(\x. (g1 * h1) x * (g2 * h2) x) = (g1 * g2) * (h1 * h2)" by (rule ext) (simp add: func_times mult_ac) from A show "(\x. (g1 * h1) x * (g2 * h2) x) \ G * H" by (subst B, intro set_times_intro) (auto intro: G.mult H.mult simp: func_times) qed (fact G.F_nontrivial) lemma smallo_iff: assumes "g1 \ G" "g2 \ G" "h1 \ H" "h2 \ H" shows "(\x. g1 x * h1 x) \ o[F](\x. g2 x * h2 x) \ g1 \ o[F](g2) \ (g1 \ \[F](g2) \ h1 \ o[F](h2))" (is "?P \ ?Q") proof (rule G.landau_cases[OF assms(1,2)]) assume "g1 \ o[F](g2)" thus ?thesis by (auto intro!: gs_dominate assms) next assume A: "g1 \ \[F](g2)" hence B: "g2 \ O[F](g1)" by (subst (asm) bigtheta_sym) (rule bigthetaD1) hence "g1 \ o[F](g2)" using assms by (auto dest: G.small_big_antisym) moreover from A have "o[F](\x. g2 x * h2 x) = o[F](\x. g1 x * h2 x)" by (intro landau_o.small.cong_bigtheta landau_theta.mult_right, subst bigtheta_sym) ultimately show ?thesis using G.nonzero[OF assms(1)] A by (auto simp add: landau_o.small.mult_cancel_left) next assume A: "g2 \ o[F](g1)" from gs_dominate[OF assms(2,1,4,3) this] have B: "g2 * h2 \ o[F](g1 * h1)" by (simp add: func_times) have "g1 \ o[F](g2)" "g1 \ \[F](g2)" using assms A by (auto dest: G.small_antisym G.small_big_antisym simp: bigomega_iff_bigo) moreover have "\?P" by (intro notI GH.small_antisym[OF _ _ B] set_times_intro) (simp_all add: func_times assms) ultimately show ?thesis by blast qed lemma bigo_iff: assumes "g1 \ G" "g2 \ G" "h1 \ H" "h2 \ H" shows "(\x. g1 x * h1 x) \ O[F](\x. g2 x * h2 x) \ g1 \ o[F](g2) \ (g1 \ \[F](g2) \ h1 \ O[F](h2))" (is "?P \ ?Q") proof (rule G.landau_cases[OF assms(1,2)]) assume "g1 \ o[F](g2)" thus ?thesis by (auto intro!: gs_dominate assms landau_o.small_imp_big) next assume A: "g2 \ o[F](g1)" hence "g1 \ O[F](g2)" using assms by (auto dest: G.small_big_antisym) moreover from gs_dominate[OF assms(2,1,4,3) A] have "g2*h2 \ o[F](g1*h1)" by (simp add: func_times) hence "g1*h1 \ O[F](g2*h2)" by (blast intro: GH.small_big_antisym assms) ultimately show ?thesis using A assms by (auto simp: func_times dest: landau_o.small_imp_big) next assume A: "g1 \ \[F](g2)" hence "g1 \ o[F](g2)" unfolding bigtheta_def using assms by (auto dest: G.small_big_antisym simp: bigomega_iff_bigo) moreover have "O[F](\x. g2 x * h2 x) = O[F](\x. g1 x * h2 x)" by (subst landau_o.big.cong_bigtheta[OF landau_theta.mult_right[OF A]]) (rule refl) ultimately show ?thesis using A G.nonzero[OF assms(2)] by (auto simp: landau_o.big.mult_cancel_left eventually_nonzero_bigtheta) qed lemma bigtheta_iff: "g1 \ G \ g2 \ G \ h1 \ H \ h2 \ H \ (\x. g1 x * h1 x) \ \[F](\x. g2 x * h2 x) \ g1 \ \[F](g2) \ h1 \ \[F](h2)" by (auto simp: bigtheta_def bigo_iff bigomega_iff_bigo intro: landau_o.small_imp_big dest: G.small_antisym G.small_big_antisym) end lemma landau_function_family_powr_closure: assumes "F \ bot" "filterlim f at_top F" shows "landau_function_family F (powr_closure f)" proof (unfold_locales; (elim powr_closureE; hypsubst)?) from assms have "eventually (\x. f x \ 1) F" using filterlim_at_top by auto hence A: "eventually (\x. f x \ 0) F" by eventually_elim simp { fix p q :: real show "(\x. f x powr p) \ o[F](\x. f x powr q) \ (\x. f x powr q) \ o[F](\x. f x powr p) \ (\x. f x powr p) \ \[F](\x. f x powr q)" by (cases p q rule: linorder_cases) (force intro!: smalloI_tendsto tendsto_neg_powr simp: powr_diff [symmetric] assms A)+ } fix p show "eventually (\x. f x powr p > 0) F" using A by simp qed (auto simp: powr_add[symmetric] powr_minus[symmetric] \F \ bot\ intro: powr_closureI) lemma landau_function_family_pair_trans: assumes "landau_function_family_pair Ftr F G f" assumes "landau_function_family_pair Ftr G H g" shows "landau_function_family_pair Ftr F (G*H) f" proof- interpret FG: landau_function_family_pair Ftr F G f by fact interpret GH: landau_function_family_pair Ftr G H g by fact show ?thesis proof (unfold_locales; (elim set_times_elim)?; (clarify)?; (unfold func_times mult.assoc[symmetric])?) fix f1 f2 g1 g2 h1 h2 assume A: "f1 \ F" "f2 \ F" "g1 \ G" "g2 \ G" "h1 \ H" "h2 \ H" "f1 \ o[Ftr](f2)" from A have "(\x. f1 x * g1 x * h1 x) \ o[Ftr](\x. f1 x * g1 x * g x)" by (intro landau_o.small.mult_left GH.g_dominates) also have "(\x. f1 x * g1 x * g x) = (\x. f1 x * (g1 x * g x))" by (simp only: mult.assoc) also from A have "... \ o[Ftr](\x. f2 x * (g2 x / g x))" by (intro FG.gs_dominate FG.H.mult FG.H.div GH.g) also from A have "(\x. inverse (h2 x)) \ o[Ftr](g)" by (intro GH.g_dominates GH.H.inverse) with GH.g A have "(\x. f2 x * (g2 x / g x)) \ o[Ftr](\x. f2 x * (g2 x * h2 x))" by (auto simp: FG.H.nonzero GH.H.nonzero divide_inverse intro!: landau_o.small.mult_left intro: landau_o.small.inverse_flip) also have "... = o[Ftr](\x. f2 x * g2 x * h2 x)" by (simp only: mult.assoc) finally show "(\x. f1 x * g1 x * h1 x) \ o[Ftr](\x. f2 x * g2 x * h2 x)" . next fix g1 h1 assume A: "g1 \ G" "h1 \ H" hence "(\x. g1 x * h1 x) \ o[Ftr](\x. g1 x * g x)" by (intro landau_o.small.mult_left GH.g_dominates) also from A have "(\x. g1 x * g x) \ o[Ftr](f)" by (intro FG.g_dominates FG.H.mult GH.g) finally show "(\x. g1 x * h1 x) \ o[Ftr](f)" . qed (simp_all add: FG.g) qed lemma landau_function_family_pair_trans_powr: assumes "landau_function_family_pair F (powr_closure g) H (\x. g x powr 1)" assumes "filterlim f at_top F" assumes "\p. (\x. g x powr p) \ o[F](f)" shows "landau_function_family_pair F (powr_closure f) (powr_closure g * H) (\x. f x powr 1)" proof (rule landau_function_family_pair_trans[OF _ assms(1)]) interpret GH: landau_function_family_pair F "powr_closure g" H "\x. g x powr 1" by fact interpret F: landau_function_family F "powr_closure f" by (rule landau_function_family_powr_closure) (rule GH.G.F_nontrivial, rule assms) show "landau_function_family_pair F (powr_closure f) (powr_closure g) (\x. f x powr 1)" proof (unfold_locales; (elim powr_closureE; hypsubst)?) show "(\x. f x powr 1) \ powr_closure f" by (rule powr_closureI) next fix p ::real note assms(3)[of p] also from assms(2) have "eventually (\x. f x \ 1) F" by (force simp: filterlim_at_top) hence "f \ \[F](\x. f x powr 1)" by (auto intro!: bigthetaI_cong elim!: eventually_mono) finally show "(\x. g x powr p) \ o[F](\x. f x powr 1)" . next fix p p1 p2 p3 :: real assume A: "(\x. f x powr p) \ o[F](\x. f x powr p1)" have p: "p < p1" proof (cases p p1 rule: linorder_cases) assume "p > p1" moreover from assms(2) have "eventually (\x. f x \ 1) F" by (force simp: filterlim_at_top) hence "eventually (\x. f x \ 0) F" by eventually_elim simp ultimately have "(\x. f x powr p1) \ o[F](\x. f x powr p)" using assms by (auto intro!: smalloI_tendsto tendsto_neg_powr simp: powr_diff [symmetric] ) from F.small_antisym[OF _ _ this A] show ?thesis by (auto simp: powr_closureI) next assume "p = p1" hence "(\x. f x powr p1) \ O[F](\x. f x powr p)" by (intro bigthetaD1) simp with F.small_big_antisym[OF _ _ A this] show ?thesis by (auto simp: powr_closureI) qed from assms(2) have f_pos: "eventually (\x. f x \ 1) F" by (force simp: filterlim_at_top) from assms have "(\x. g x powr ((p2 - p3)/(p1 - p))) \ o[F](f)" by simp from smallo_powr[OF this, of "p1 - p"] p have "(\x. g x powr (p2 - p3)) \ o[F](\x. \f x\ powr (p1 - p))" by (simp add: powr_powr) hence "(\x. \f x\ powr p * g x powr p2) \ o[F](\x. \f x\ powr p1 * g x powr p3)" (is ?P) using GH.G.nonzero[OF GH.g] F.nonzero[OF powr_closureI] by (simp add: powr_diff landau_o.small.divide_eq1 landau_o.small.divide_eq2 mult.commute) also have "?P \ (\x. f x powr p * g x powr p2) \ o[F](\x. f x powr p1 * g x powr p3)" using f_pos by (intro landau_o.small.cong_ex) (auto elim!: eventually_mono) finally show "(\x. f x powr p * g x powr p2) \ o[F](\x. f x powr p1 * g x powr p3)" . qed qed definition dominates :: "'a filter \ ('a \ real) \ ('a \ real) \ bool" where "dominates F f g = (\p. (\x. g x powr p) \ o[F](f))" lemma dominates_trans: assumes "eventually (\x. g x > 0) F" assumes "dominates F f g" "dominates F g h" shows "dominates F f h" unfolding dominates_def proof fix p :: real from assms(3) have "(\x. h x powr p) \ o[F](g)" unfolding dominates_def by simp also from assms(1) have "g \ \[F](\x. g x powr 1)" by (intro bigthetaI_cong) (auto elim!: eventually_mono) also from assms(2) have "(\x. g x powr 1) \ o[F](f)" unfolding dominates_def by simp finally show "(\x. h x powr p) \ o[F](f)" . qed fun landau_dominating_chain where "landau_dominating_chain F (f # g # gs) \ dominates F f g \ landau_dominating_chain F (g # gs)" | "landau_dominating_chain F [f] \ (\x. 1) \ o[F](f)" | "landau_dominating_chain F [] \ True" primrec landau_dominating_chain' where "landau_dominating_chain' F [] \ True" | "landau_dominating_chain' F (f # gs) \ landau_function_family_pair F (powr_closure f) (prod_list (map powr_closure gs)) (\x. f x powr 1) \ landau_dominating_chain' F gs" primrec nonneg_list where "nonneg_list [] \ True" | "nonneg_list (x#xs) \ x > 0 \ (x = 0 \ nonneg_list xs)" primrec pos_list where "pos_list [] \ False" | "pos_list (x#xs) \ x > 0 \ (x = 0 \ pos_list xs)" lemma dominating_chain_imp_dominating_chain': "Ftr \ bot \ (\g. g \ set gs \ filterlim g at_top Ftr) \ landau_dominating_chain Ftr gs \ landau_dominating_chain' Ftr gs" proof (induction gs rule: landau_dominating_chain.induct) case (1 F f g gs) from 1 show ?case by (auto intro!: landau_function_family_pair_trans_powr simp add: dominates_def) next case (2 F f) then interpret F: landau_function_family F "powr_closure f" by (intro landau_function_family_powr_closure) simp_all from 2 have "eventually (\x. f x \ 1) F" by (force simp: filterlim_at_top) hence "o[F](\x. f x powr 1) = o[F](\x. f x)" by (intro landau_o.small.cong) (auto elim!: eventually_mono) with 2 have "landau_function_family_pair F (powr_closure f) {\_. 1} (\x. f x powr 1)" by unfold_locales (auto intro: powr_closureI) thus ?case by (simp add: one_fun_def) next case 3 then show ?case by simp qed locale landau_function_family_chain = fixes F :: "'b filter" fixes gs :: "'a list" fixes get_param :: "'a \ real" fixes get_fun :: "'a \ ('b \ real)" assumes F_nontrivial: "F \ bot" assumes gs_pos: "g \ set (map get_fun gs) \ filterlim g at_top F" assumes dominating_chain: "landau_dominating_chain F (map get_fun gs)" begin lemma dominating_chain': "landau_dominating_chain' F (map get_fun gs)" by (intro dominating_chain_imp_dominating_chain' gs_pos dominating_chain F_nontrivial) lemma gs_powr_0_eq_one: "eventually (\x. (\g\gs. get_fun g x powr 0) = 1) F" using gs_pos proof (induction gs) case (Cons g gs) from Cons have "eventually (\x. get_fun g x > 0) F" by (auto simp: filterlim_at_top_dense) moreover from Cons have "eventually (\x. (\g\gs. get_fun g x powr 0) = 1) F" by simp ultimately show ?case by eventually_elim simp qed simp_all lemma listmap_gs_in_listmap: "(\x. \g\fs. h g x powr p g) \ prod_list (map powr_closure (map h fs))" proof- have "(\x. \g\fs. h g x powr p g) = (\g\fs. (\x. h g x powr p g))" by (rule ext, induction fs) simp_all also have "... \ prod_list (map powr_closure (map h fs))" apply (induction fs) apply (simp add: fun_eq_iff) apply (simp only: list.map prod_list.Cons, rule set_times_intro) apply simp_all done finally show ?thesis . qed lemma smallo_iff: "(\_. 1) \ o[F](\x. \g\gs. get_fun g x powr get_param g) \ pos_list (map get_param gs)" proof- have "((\_. 1) \ o[F](\x. \g\gs. get_fun g x powr get_param g)) \ ((\x. \g\gs. get_fun g x powr 0) \ o[F](\x. \g\gs. get_fun g x powr get_param g))" by (rule sym, intro landau_o.small.in_cong gs_powr_0_eq_one) also from gs_pos dominating_chain' have "... \ pos_list (map get_param gs)" proof (induction gs) case Nil have "(\x::'b. 1::real) \ o[F](\x. 1)" using F_nontrivial by (auto dest!: landau_o.small_big_asymmetric) thus ?case by simp next case (Cons g gs) then interpret G: landau_function_family_pair F "powr_closure (get_fun g)" "prod_list (map powr_closure (map get_fun gs))" "\x. get_fun g x powr 1" by simp from Cons show ?case using listmap_gs_in_listmap[of get_fun _ gs] F_nontrivial by (simp_all add: G.smallo_iff listmap_gs_in_listmap powr_smallo_iff powr_bigtheta_iff del: powr_zero_eq_one) qed finally show ?thesis . qed lemma bigo_iff: "(\_. 1) \ O[F](\x. \g\gs. get_fun g x powr get_param g) \ nonneg_list (map get_param gs)" proof- have "((\_. 1) \ O[F](\x. \g\gs. get_fun g x powr get_param g)) \ ((\x. \g\gs. get_fun g x powr 0) \ O[F](\x. \g\gs. get_fun g x powr get_param g))" by (rule sym, intro landau_o.big.in_cong gs_powr_0_eq_one) also from gs_pos dominating_chain' have "... \ nonneg_list (map get_param gs)" proof (induction gs) case Nil then show ?case by (simp add: func_one) next case (Cons g gs) then interpret G: landau_function_family_pair F "powr_closure (get_fun g)" "prod_list (map powr_closure (map get_fun gs))" "\x. get_fun g x powr 1" by simp from Cons show ?case using listmap_gs_in_listmap[of get_fun _ gs] F_nontrivial by (simp_all add: G.bigo_iff listmap_gs_in_listmap powr_smallo_iff powr_bigtheta_iff del: powr_zero_eq_one) qed finally show ?thesis . qed lemma bigtheta_iff: "(\_. 1) \ \[F](\x. \g\gs. get_fun g x powr get_param g) \ list_all ((=) 0) (map get_param gs)" proof- have "((\_. 1) \ \[F](\x. \g\gs. get_fun g x powr get_param g)) \ ((\x. \g\gs. get_fun g x powr 0) \ \[F](\x. \g\gs. get_fun g x powr get_param g))" by (rule sym, intro landau_theta.in_cong gs_powr_0_eq_one) also from gs_pos dominating_chain' have "... \ list_all ((=) 0) (map get_param gs)" proof (induction gs) case Nil then show ?case by (simp add: func_one) next case (Cons g gs) then interpret G: landau_function_family_pair F "powr_closure (get_fun g)" "prod_list (map powr_closure (map get_fun gs))" "\x. get_fun g x powr 1" by simp from Cons show ?case using listmap_gs_in_listmap[of get_fun _ gs] F_nontrivial by (simp_all add: G.bigtheta_iff listmap_gs_in_listmap powr_smallo_iff powr_bigtheta_iff del: powr_zero_eq_one) qed finally show ?thesis . qed end lemma fun_chain_at_top_at_top: assumes "filterlim (f :: ('a::order) \ 'a) at_top at_top" shows "filterlim (f ^^ n) at_top at_top" by (induction n) (auto intro: filterlim_ident filterlim_compose[OF assms]) lemma const_smallo_ln_chain: "(\_. 1) \ o((ln::real\real)^^n)" proof (intro smalloI_tendsto) show "((\x::real. 1 / (ln^^n) x) \ 0) at_top" by (rule tendsto_divide_0 tendsto_const filterlim_at_top_imp_at_infinity fun_chain_at_top_at_top ln_at_top)+ next from fun_chain_at_top_at_top[OF ln_at_top, of n] have "eventually (\x::real. (ln^^n) x > 0) at_top" by (simp add: filterlim_at_top_dense) thus "eventually (\x::real. (ln^^n) x \ 0) at_top" by eventually_elim simp_all qed lemma ln_fun_in_smallo_fun: assumes "filterlim f at_top at_top" shows "(\x. ln (f x) powr p :: real) \ o(f)" proof (rule smalloI_tendsto) have "((\x. ln x powr p / x powr 1) \ 0) at_top" by (rule tendsto_ln_powr_over_powr') simp moreover have "eventually (\x. ln x powr p / x powr 1 = ln x powr p / x) at_top" using eventually_gt_at_top[of "0::real"] by eventually_elim simp ultimately have "((\x. ln x powr p / x) \ 0) at_top" by (subst (asm) tendsto_cong) from this assms show "((\x. ln (f x) powr p / f x) \ 0) at_top" by (rule filterlim_compose) from assms have "eventually (\x. f x \ 1) at_top" by (simp add: filterlim_at_top) thus "eventually (\x. f x \ 0) at_top" by eventually_elim simp qed lemma ln_chain_dominates: "m > n \ dominates at_top ((ln::real \ real)^^n) (ln^^m)" proof (erule less_Suc_induct) fix n show "dominates at_top ((ln::real\real)^^n) (ln^^(Suc n))" unfolding dominates_def by (force intro: ln_fun_in_smallo_fun fun_chain_at_top_at_top ln_at_top) next fix k m n assume A: "dominates at_top ((ln::real \ real)^^k) (ln^^m)" "dominates at_top ((ln::real \ real)^^m) (ln^^n)" from fun_chain_at_top_at_top[OF ln_at_top, of m] have "eventually (\x::real. (ln^^m) x > 0) at_top" by (simp add: filterlim_at_top_dense) from this A show "dominates at_top ((ln::real \ real)^^k) ((ln::real \ real)^^n)" by (rule dominates_trans) qed datatype primfun = LnChain nat instantiation primfun :: linorder begin fun less_eq_primfun :: "primfun \ primfun \ bool" where "LnChain x \ LnChain y \ x \ y" fun less_primfun :: "primfun \ primfun \ bool" where "LnChain x < LnChain y \ x < y" instance proof (standard, goal_cases) case (1 x y) show ?case by (induction x y rule: less_eq_primfun.induct) auto next case (2 x) show ?case by (cases x) auto next case (3 x y z) thus ?case by (induction x y rule: less_eq_primfun.induct, cases z) auto next case (4 x y) thus ?case by (induction x y rule: less_eq_primfun.induct) auto next case (5 x y) thus ?case by (induction x y rule: less_eq_primfun.induct) auto qed end fun eval_primfun' :: "_ \ _ \ real" where "eval_primfun' (LnChain n) = (\x. (ln^^n) x)" fun eval_primfun :: "_ \ _ \ real" where "eval_primfun (f, e) = (\x. eval_primfun' f x powr e)" lemma eval_primfun_altdef: "eval_primfun f x = eval_primfun' (fst f) x powr snd f" by (cases f) simp fun merge_primfun where "merge_primfun (x::primfun, a) (y, b) = (x, a + b)" fun inverse_primfun where "inverse_primfun (x::primfun, a) = (x, -a)" fun powr_primfun where "powr_primfun (x::primfun, a) e = (x, e*a)" lemma primfun_cases: assumes "(\n e. P (LnChain n, e))" shows "P x" proof (cases x, hypsubst) fix a b show "P (a, b)" by (cases a; hypsubst, rule assms) qed lemma eval_primfun'_at_top: "filterlim (eval_primfun' f) at_top at_top" by (cases f) (auto intro!: fun_chain_at_top_at_top ln_at_top) lemma primfun_dominates: "f < g \ dominates at_top (eval_primfun' f) (eval_primfun' g)" by (elim less_primfun.elims; hypsubst) (simp_all add: ln_chain_dominates) lemma eval_primfun_pos: "eventually (\x::real. eval_primfun f x > 0) at_top" proof (cases f, hypsubst) fix f e from eval_primfun'_at_top have "eventually (\x. eval_primfun' f x > 0) at_top" by (auto simp: filterlim_at_top_dense) thus "eventually (\x::real. eval_primfun (f,e) x > 0) at_top" by eventually_elim simp qed lemma eventually_nonneg_primfun: "eventually_nonneg at_top (eval_primfun f)" unfolding eventually_nonneg_def using eval_primfun_pos[of f] by eventually_elim simp lemma eval_primfun_nonzero: "eventually (\x. eval_primfun f x \ 0) at_top" using eval_primfun_pos[of f] by eventually_elim simp lemma eval_merge_primfun: "fst f = fst g \ eval_primfun (merge_primfun f g) x = eval_primfun f x * eval_primfun g x" by (induction f g rule: merge_primfun.induct) (simp_all add: powr_add) lemma eval_inverse_primfun: "eval_primfun (inverse_primfun f) x = inverse (eval_primfun f x)" by (induction f rule: inverse_primfun.induct) (simp_all add: powr_minus) lemma eval_powr_primfun: "eval_primfun (powr_primfun f e) x = eval_primfun f x powr e" by (induction f e rule: powr_primfun.induct) (simp_all add: powr_powr mult.commute) definition eval_primfuns where "eval_primfuns fs x = (\f\fs. eval_primfun f x)" lemma eval_primfuns_pos: "eventually (\x. eval_primfuns fs x > 0) at_top" proof- have prod_list_pos: "(\x::_::linordered_semidom. x \ set xs \ x > 0) \ prod_list xs > 0" for xs :: "real list" by (induction xs) auto have "eventually (\x. \f\set fs. eval_primfun f x > 0) at_top" by (intro eventually_ball_finite ballI eval_primfun_pos finite_set) thus ?thesis unfolding eval_primfuns_def by eventually_elim (rule prod_list_pos, auto) qed lemma eval_primfuns_nonzero: "eventually (\x. eval_primfuns fs x \ 0) at_top" using eval_primfuns_pos[of fs] by eventually_elim simp subsection \Reification\ definition LANDAU_PROD' where "LANDAU_PROD' L c f = L(\x. c * f x)" definition LANDAU_PROD where "LANDAU_PROD L c1 c2 fs \ (\_. c1) \ L(\x. c2 * eval_primfuns fs x)" definition BIGTHETA_CONST' where "BIGTHETA_CONST' c = \(\x. c)" definition BIGTHETA_CONST where "BIGTHETA_CONST c A = set_mult \(\_. c) A" definition BIGTHETA_FUN where "BIGTHETA_FUN f = \(f)" lemma BIGTHETA_CONST'_tag: "\(\x. c) = BIGTHETA_CONST' c" using BIGTHETA_CONST'_def .. lemma BIGTHETA_CONST_tag: "\(f) = BIGTHETA_CONST 1 \(f)" by (simp add: BIGTHETA_CONST_def bigtheta_mult_eq_set_mult[symmetric]) lemma BIGTHETA_FUN_tag: "\(f) = BIGTHETA_FUN f" by (simp add: BIGTHETA_FUN_def) lemma set_mult_is_times: "set_mult A B = A * B" unfolding set_mult_def set_times_def func_times by blast lemma set_powr_mult: assumes "eventually_nonneg F f" and "eventually_nonneg F g" shows "\[F](\x. (f x * g x :: real) powr p) = set_mult (\[F](\x. f x powr p)) (\[F](\x. g x powr p))" proof- from assms have "eventually (\x. f x \ 0) F" "eventually (\x. g x \ 0) F" by (simp_all add: eventually_nonneg_def) hence "eventually (\x. (f x * g x :: real) powr p = f x powr p * g x powr p) F" by eventually_elim (simp add: powr_mult) hence "\[F](\x. (f x * g x :: real) powr p) = \[F](\x. f x powr p * g x powr p)" by (rule landau_theta.cong) also have "... = set_mult (\[F](\x. f x powr p)) (\[F](\x. g x powr p))" by (simp add: bigtheta_mult_eq_set_mult) finally show ?thesis . qed lemma eventually_nonneg_bigtheta_pow_realpow: "\(\x. eval_primfun f x ^ e) = \(\x. eval_primfun f x powr real e)" using eval_primfun_pos[of f] by (auto intro!: landau_theta.cong elim!: eventually_mono simp: powr_realpow) lemma BIGTHETA_CONST_fold: "BIGTHETA_CONST (c::real) (BIGTHETA_CONST d A) = BIGTHETA_CONST (c*d) A" "bigtheta_pow at_top (BIGTHETA_CONST c \(eval_primfun pf)) k = BIGTHETA_CONST (c ^ k) \(\x. eval_primfun pf x powr k)" "set_inverse (BIGTHETA_CONST c \(f)) = BIGTHETA_CONST (inverse c) \(\x. inverse (f x))" "set_mult (BIGTHETA_CONST c \(f)) (BIGTHETA_CONST d \(g)) = BIGTHETA_CONST (c*d) \(\x. f x*g x)" "BIGTHETA_CONST' (c::real) = BIGTHETA_CONST c \(\_. 1)" "BIGTHETA_FUN (f::real\real) = BIGTHETA_CONST 1 \(f)" apply (simp add: BIGTHETA_CONST_def set_mult_is_times bigtheta_mult_eq_set_mult mult_ac) apply (simp only: BIGTHETA_CONST_def bigtheta_mult_eq_set_mult[symmetric] bigtheta_pow_eq_set_pow[symmetric] power_mult_distrib mult_ac) apply (simp add: bigtheta_mult_eq_set_mult eventually_nonneg_bigtheta_pow_realpow) by (simp_all add: BIGTHETA_CONST_def BIGTHETA_CONST'_def BIGTHETA_FUN_def bigtheta_mult_eq_set_mult[symmetric] set_mult_is_times[symmetric] bigtheta_pow_eq_set_pow[symmetric] bigtheta_inverse_eq_set_inverse[symmetric] mult_ac power_mult_distrib) lemma fold_fun_chain: "g x = (g ^^ 1) x" "(g ^^ m) ((g ^^ n) x) = (g ^^ (m+n)) x" by (simp_all add: funpow_add) -lemma reify_ln_chain_1: +lemma reify_ln_chain1: "\(\x. (ln ^^ n) x) = \(eval_primfun (LnChain n, 1))" proof (intro landau_theta.cong) have "filterlim ((ln :: real \ real) ^^ n) at_top at_top" by (intro fun_chain_at_top_at_top ln_at_top) hence "eventually (\x::real. (ln ^^ n) x > 0) at_top" using filterlim_at_top_dense by auto thus "eventually (\x. (ln ^^ n) x = eval_primfun (LnChain n, 1) x) at_top" by eventually_elim simp qed -lemma reify_monom_1: +lemma reify_monom1: "\(\x::real. x) = \(eval_primfun (LnChain 0, 1))" proof (intro landau_theta.cong) from eventually_gt_at_top[of "0::real"] show "eventually (\x. x = eval_primfun (LnChain 0, 1) x) at_top" by eventually_elim simp qed lemma reify_monom_pow: "\(\x::real. x ^ e) = \(eval_primfun (LnChain 0, real e))" proof- have "\(eval_primfun (LnChain 0, real e)) = \(\x. x powr (real e))" by simp also have "eventually (\x. x powr (real e) = x ^ e) at_top" using eventually_gt_at_top[of 0] by eventually_elim (simp add: powr_realpow) hence "\(\x. x powr (real e)) = \(\x. x ^ e)" by (rule landau_theta.cong) finally show ?thesis .. qed lemma reify_monom_powr: "\(\x::real. x powr e) = \(eval_primfun (LnChain 0, e))" by (rule landau_theta.cong) simp -lemmas reify_monom = reify_monom_1 reify_monom_pow reify_monom_powr +lemmas reify_monom = reify_monom1 reify_monom_pow reify_monom_powr lemma reify_ln_chain_pow: "\(\x. (ln ^^ n) x ^ e) = \(eval_primfun (LnChain n, real e))" proof- have "\(eval_primfun (LnChain n, real e)) = \(\x. (ln ^^ n) x powr (real e))" by simp also have "eventually (\x::real. (ln ^^ n) x > 0) at_top" using fun_chain_at_top_at_top[OF ln_at_top] unfolding filterlim_at_top_dense by blast hence "eventually (\x. (ln ^^ n) x powr (real e) = (ln ^^ n) x ^ e) at_top" by eventually_elim (subst powr_realpow, auto) hence "\(\x. (ln ^^ n) x powr (real e)) = \(\x. (ln ^^ n) x^e)" by (rule landau_theta.cong) finally show ?thesis .. qed lemma reify_ln_chain_powr: "\(\x. (ln ^^ n) x powr e) = \(eval_primfun (LnChain n, e))" by (intro landau_theta.cong) simp -lemmas reify_ln_chain = reify_ln_chain_1 reify_ln_chain_pow reify_ln_chain_powr +lemmas reify_ln_chain = reify_ln_chain1 reify_ln_chain_pow reify_ln_chain_powr lemma numeral_power_Suc: "numeral n ^ Suc a = numeral n * numeral n ^ a" by (rule power.simps) lemmas landau_product_preprocess = one_add_one one_plus_numeral numeral_plus_one arith_simps numeral_power_Suc power_0 fold_fun_chain[where g = ln] reify_ln_chain reify_monom lemma LANDAU_PROD'_fold: "BIGTHETA_CONST e \(\_. d) = BIGTHETA_CONST (e*d) \(eval_primfuns [])" "LANDAU_PROD' c (\_. 1) = LANDAU_PROD' c (eval_primfuns [])" "eval_primfun f = eval_primfuns [f]" "eval_primfuns fs x * eval_primfuns gs x = eval_primfuns (fs @ gs) x" apply (simp only: BIGTHETA_CONST_def set_mult_is_times eval_primfuns_def[abs_def] bigtheta_mult_eq) apply (simp add: bigtheta_mult_eq[symmetric]) by (simp_all add: eval_primfuns_def[abs_def] BIGTHETA_CONST_def) lemma inverse_prod_list_field: "prod_list (map (\x. inverse (f x)) xs) = inverse (prod_list (map f xs :: _ :: field list))" by (induction xs) simp_all lemma landau_prod_meta_cong: assumes "landau_symbol L L' Lr" assumes "\(f) \ BIGTHETA_CONST c1 (\(eval_primfuns fs))" assumes "\(g) \ BIGTHETA_CONST c2 (\(eval_primfuns gs))" shows "f \ L at_top (g) \ LANDAU_PROD (L at_top) c1 c2 (map inverse_primfun fs @ gs)" proof- interpret landau_symbol L L' Lr by fact have "f \ L at_top (g) \ (\x. c1 * eval_primfuns fs x) \ L at_top (\x. c2 * eval_primfuns gs x)" using assms(2,3)[symmetric] unfolding BIGTHETA_CONST_def by (intro cong_ex_bigtheta) (simp_all add: bigtheta_mult_eq_set_mult[symmetric]) also have "... \ (\x. c1) \ L at_top (\x. c2 * eval_primfuns gs x / eval_primfuns fs x)" by (simp_all add: eval_primfuns_nonzero divide_eq1) finally show "f \ L at_top (g) \ LANDAU_PROD (L at_top) c1 c2 (map inverse_primfun fs @ gs)" by (simp add: LANDAU_PROD_def eval_primfuns_def eval_inverse_primfun divide_inverse o_def inverse_prod_list_field mult_ac) qed fun pos_primfun_list where "pos_primfun_list [] \ False" | "pos_primfun_list ((_,x)#xs) \ x > 0 \ (x = 0 \ pos_primfun_list xs)" fun nonneg_primfun_list where "nonneg_primfun_list [] \ True" | "nonneg_primfun_list ((_,x)#xs) \ x > 0 \ (x = 0 \ nonneg_primfun_list xs)" fun iszero_primfun_list where "iszero_primfun_list [] \ True" | "iszero_primfun_list ((_,x)#xs) \ x = 0 \ iszero_primfun_list xs" definition "group_primfuns \ groupsort.group_sort fst merge_primfun" lemma list_ConsCons_induct: assumes "P []" "\x. P [x]" "\x y xs. P (y#xs) \ P (x#y#xs)" shows "P xs" proof (induction xs rule: length_induct) case (1 xs) show ?case proof (cases xs) case (Cons x xs') note A = this from assms 1 show ?thesis proof (cases xs') case (Cons y xs'') with 1 A have "P (y#xs'')" by simp with Cons A assms show ?thesis by simp qed (simp add: assms A) qed (simp add: assms) qed lemma landau_function_family_chain_primfuns: assumes "sorted (map fst fs)" assumes "distinct (map fst fs)" shows "landau_function_family_chain at_top fs (eval_primfun' o fst)" proof (standard, goal_cases) case 3 from assms show ?case proof (induction fs rule: list_ConsCons_induct) case (2 g) from eval_primfun'_at_top[of "fst g"] have "eval_primfun' (fst g) \ \(\_. 1)" by (intro smallomegaI_filterlim_at_infinity filterlim_at_top_imp_at_infinity) simp thus ?case by (simp add: smallomega_iff_smallo) next case (3 f g gs) thus ?case by (auto simp: primfun_dominates) qed simp qed (auto simp: eval_primfun'_at_top) lemma (in monoid_mult) fold_plus_prod_list_rev: "fold times xs = times (prod_list (rev xs))" proof fix x have "fold times xs x = prod_list (rev xs @ [x])" by (simp add: foldr_conv_fold prod_list.eq_foldr) also have "\ = prod_list (rev xs) * x" by simp finally show "fold times xs x = prod_list (rev xs) * x" . qed interpretation groupsort_primfun: groupsort fst merge_primfun eval_primfuns proof (standard, goal_cases) case (1 x y) thus ?case by (induction x y rule: merge_primfun.induct) simp_all next case (2 fs gs) show ?case proof fix x have "eval_primfuns fs x = fold (*) (map (\f. eval_primfun f x) fs) 1" unfolding eval_primfuns_def by (simp add: fold_plus_prod_list_rev) also have "fold (*) (map (\f. eval_primfun f x) fs) = fold (*) (map (\f. eval_primfun f x) gs)" using 2 by (intro fold_multiset_equiv ext) auto also have "... 1 = eval_primfuns gs x" unfolding eval_primfuns_def by (simp add: fold_plus_prod_list_rev) finally show "eval_primfuns fs x = eval_primfuns gs x" . qed qed (auto simp: fun_eq_iff eval_merge_primfun eval_primfuns_def) lemma nonneg_primfun_list_iff: "nonneg_primfun_list fs = nonneg_list (map snd fs)" by (induction fs rule: nonneg_primfun_list.induct) simp_all lemma pos_primfun_list_iff: "pos_primfun_list fs = pos_list (map snd fs)" by (induction fs rule: pos_primfun_list.induct) simp_all lemma iszero_primfun_list_iff: "iszero_primfun_list fs = list_all ((=) 0) (map snd fs)" by (induction fs rule: iszero_primfun_list.induct) simp_all lemma landau_primfuns_iff: "((\_. 1) \ O(eval_primfuns fs)) = nonneg_primfun_list (group_primfuns fs)" (is "?A") "((\_. 1) \ o(eval_primfuns fs)) = pos_primfun_list (group_primfuns fs)" (is "?B") "((\_. 1) \ \(eval_primfuns fs)) = iszero_primfun_list (group_primfuns fs)" (is "?C") proof- interpret landau_function_family_chain at_top "group_primfuns fs" snd "eval_primfun' o fst" by (rule landau_function_family_chain_primfuns) (simp_all add: group_primfuns_def groupsort_primfun.sorted_group_sort groupsort_primfun.distinct_group_sort) have "(\_. 1) \ O(eval_primfuns fs) \ (\_. 1) \ O(eval_primfuns (group_primfuns fs))" by (simp_all add: groupsort_primfun.g_group_sort group_primfuns_def) also have "... \ nonneg_list (map snd (group_primfuns fs))" using bigo_iff by (simp add: eval_primfuns_def[abs_def] eval_primfun_altdef) finally show ?A by (simp add: nonneg_primfun_list_iff) have "(\_. 1) \ o(eval_primfuns fs) \ (\_. 1) \ o(eval_primfuns (group_primfuns fs))" by (simp_all add: groupsort_primfun.g_group_sort group_primfuns_def) also have "... \ pos_list (map snd (group_primfuns fs))" using smallo_iff by (simp add: eval_primfuns_def[abs_def] eval_primfun_altdef) finally show ?B by (simp add: pos_primfun_list_iff) have "(\_. 1) \ \(eval_primfuns fs) \ (\_. 1) \ \(eval_primfuns (group_primfuns fs))" by (simp_all add: groupsort_primfun.g_group_sort group_primfuns_def) also have "... \ list_all ((=) 0) (map snd (group_primfuns fs))" using bigtheta_iff by (simp add: eval_primfuns_def[abs_def] eval_primfun_altdef) finally show ?C by (simp add: iszero_primfun_list_iff) qed lemma LANDAU_PROD_bigo_iff: "LANDAU_PROD (bigo at_top) c1 c2 fs \ c1 = 0 \ (c2 \ 0 \ nonneg_primfun_list (group_primfuns fs))" unfolding LANDAU_PROD_def by (cases "c1 = 0", simp, cases "c2 = 0", simp) (simp_all add: landau_primfuns_iff) lemma LANDAU_PROD_smallo_iff: "LANDAU_PROD (smallo at_top) c1 c2 fs \ c1 = 0 \ (c2 \ 0 \ pos_primfun_list (group_primfuns fs))" unfolding LANDAU_PROD_def by (cases "c1 = 0", simp, cases "c2 = 0", simp) (simp_all add: landau_primfuns_iff) lemma LANDAU_PROD_bigtheta_iff: "LANDAU_PROD (bigtheta at_top) c1 c2 fs \ (c1 = 0 \ c2 = 0) \ (c1 \ 0 \ c2 \ 0 \ iszero_primfun_list (group_primfuns fs))" proof- have A: "\P x. (x = 0 \ P) \ (x \ 0 \ P) \ P" by blast { assume "eventually (\x. eval_primfuns fs x = 0) at_top" with eval_primfuns_nonzero[of fs] have "eventually (\x::real. False) at_top" by eventually_elim simp hence False by simp } note B = this show ?thesis by (rule A[of c1, case_product A[of c2]]) (insert B, auto simp: LANDAU_PROD_def landau_primfuns_iff) qed lemmas LANDAU_PROD_iff = LANDAU_PROD_bigo_iff LANDAU_PROD_smallo_iff LANDAU_PROD_bigtheta_iff lemmas landau_real_prod_simps [simp] = groupsort_primfun.group_part_def group_primfuns_def groupsort_primfun.group_sort.simps groupsort_primfun.group_part_aux.simps pos_primfun_list.simps nonneg_primfun_list.simps iszero_primfun_list.simps end diff --git a/thys/Laplace_Transform/Laplace_Transform_Library.thy b/thys/Laplace_Transform/Laplace_Transform_Library.thy --- a/thys/Laplace_Transform/Laplace_Transform_Library.thy +++ b/thys/Laplace_Transform/Laplace_Transform_Library.thy @@ -1,264 +1,264 @@ theory Laplace_Transform_Library imports "HOL-Analysis.Analysis" begin section \References\ text \ Much of this formalization is based on Schiff's textbook @{cite "Schiff2013"}. Parts of this formalization are inspired by the HOL-Light formalization (@{cite "Taqdees2013"}, @{cite "Rashid2017"}, @{cite "Rashid2018"}), but stated more generally for piecewise continuous (instead of piecewise continuously differentiable) functions. \ section \Library Additions\ subsection \Derivatives\ lemma DERIV_compose_FDERIV:\\TODO: generalize and move from HOL-ODE\ assumes "DERIV f (g x) :> f'" assumes "(g has_derivative g') (at x within s)" shows "((\x. f (g x)) has_derivative (\x. g' x * f')) (at x within s)" using assms has_derivative_compose[of g g' x s f "(*) f'"] by (auto simp: has_field_derivative_def ac_simps) lemmas has_derivative_sin[derivative_intros] = DERIV_sin[THEN DERIV_compose_FDERIV] and has_derivative_cos[derivative_intros] = DERIV_cos[THEN DERIV_compose_FDERIV] and has_derivative_exp[derivative_intros] = DERIV_exp[THEN DERIV_compose_FDERIV] subsection \Integrals\ lemma negligible_real_ivlI: fixes a b::real assumes "a \ b" shows "negligible {a .. b}" proof - from assms have "{a .. b} = {a} \ {a .. b} = {}" by auto then show ?thesis by auto qed lemma absolutely_integrable_on_combine: fixes f :: "real \ 'a::euclidean_space" assumes "f absolutely_integrable_on {a..c}" and "f absolutely_integrable_on {c..b}" and "a \ c" and "c \ b" shows "f absolutely_integrable_on {a..b}" using assms unfolding absolutely_integrable_on_def integrable_on_def by (auto intro!: has_integral_combine) lemma dominated_convergence_at_top: fixes f :: "real \ 'n::euclidean_space \ 'm::euclidean_space" assumes f: "\k. (f k) integrable_on s" and h: "h integrable_on s" and le: "\k x. x \ s \ norm (f k x) \ h x" and conv: "\x \ s. ((\k. f k x) \ g x) at_top" shows "g integrable_on s" "((\k. integral s (f k)) \ integral s g) at_top" proof - have 3: "set_integrable lebesgue s h" unfolding absolutely_integrable_on_def proof show "(\x. norm (h x)) integrable_on s" proof (intro integrable_spike_finite[OF _ _ h, where S="{}"] ballI) fix x assume "x \ s - {}" then show "norm (h x) = h x" using order_trans[OF norm_ge_zero le[of x]] by auto qed auto qed fact have 2: "set_borel_measurable lebesgue s (f k)" for k using f[of k] using has_integral_implies_lebesgue_measurable[of "f k"] by (auto intro: simp: integrable_on_def set_borel_measurable_def) have conv': "\x \ s. ((\k. f k x) \ g x) sequentially" using conv filterlim_filtermap filterlim_compose filterlim_real_sequentially by blast from 2 have 1: "set_borel_measurable lebesgue s g" unfolding set_borel_measurable_def by (rule borel_measurable_LIMSEQ_metric) (use conv' in \auto split: split_indicator\) have 4: "AE x in lebesgue. ((\i. indicator s x *\<^sub>R f i x) \ indicator s x *\<^sub>R g x) at_top" "\\<^sub>F i in at_top. AE x in lebesgue. norm (indicator s x *\<^sub>R f i x) \ indicator s x *\<^sub>R h x" using conv le by (auto intro!: always_eventually split: split_indicator) note 1 2 3 4 note * = this[unfolded set_borel_measurable_def set_integrable_def] have g: "g absolutely_integrable_on s" unfolding set_integrable_def by (rule integrable_dominated_convergence_at_top[OF *]) then show "g integrable_on s" by (auto simp: absolutely_integrable_on_def) have "((\k. (LINT x:s|lebesgue. f k x)) \ (LINT x:s|lebesgue. g x)) at_top" unfolding set_lebesgue_integral_def using * by (rule integral_dominated_convergence_at_top) then show "((\k. integral s (f k)) \ integral s g) at_top" using g absolutely_integrable_integrable_bound[OF le f h] by (subst (asm) (1 2) set_lebesgue_integral_eq_integral) auto qed lemma has_integral_dominated_convergence_at_top: fixes f :: "real \ 'n::euclidean_space \ 'm::euclidean_space" assumes "\k. (f k has_integral y k) s" "h integrable_on s" "\k x. x\s \ norm (f k x) \ h x" "\x\s. ((\k. f k x) \ g x) at_top" and x: "(y \ x) at_top" shows "(g has_integral x) s" proof - have int_f: "\k. (f k) integrable_on s" using assms by (auto simp: integrable_on_def) have "(g has_integral (integral s g)) s" by (intro integrable_integral dominated_convergence_at_top[OF int_f assms(2)]) fact+ moreover have "integral s g = x" proof (rule tendsto_unique) show "((\i. integral s (f i)) \ x) at_top" using integral_unique[OF assms(1)] x by simp show "((\i. integral s (f i)) \ integral s g) at_top" by (intro dominated_convergence_at_top[OF int_f assms(2)]) fact+ qed simp ultimately show ?thesis by simp qed lemma integral_indicator_eq_restriction: fixes f::"'a::euclidean_space \ 'b::banach" assumes f: "f integrable_on R" and "R \ S" shows "integral S (\x. indicator R x *\<^sub>R f x) = integral R f" proof - let ?f = "\x. indicator R x *\<^sub>R f x" have "?f integrable_on R" using f negligible_empty by (rule integrable_spike) auto from integrable_integral[OF this] have "(?f has_integral integral R ?f) S" by (rule has_integral_on_superset) (use \R \ S\ in \auto simp: indicator_def\) also have "integral R ?f = integral R f" using negligible_empty by (rule integral_spike) auto finally show ?thesis by blast qed lemma improper_integral_at_top: fixes f::"real \ 'a::euclidean_space" assumes "f absolutely_integrable_on {a..}" shows "((\x. integral {a..x} f) \ integral {a..} f) at_top" proof - let ?f = "\(k::real) (t::real). indicator {a..k} t *\<^sub>R f t" have f: "f integrable_on {a..k}" for k using set_lebesgue_integral_eq_integral(1)[OF assms] by (rule integrable_on_subinterval) simp from this negligible_empty have "?f k integrable_on {a..k}" for k by (rule integrable_spike) auto from this have "?f k integrable_on {a..}" for k by (rule integrable_on_superset) auto moreover have "(\x. norm (f x)) integrable_on {a..}" using assms by (simp add: absolutely_integrable_on_def) moreover note _ moreover have "\\<^sub>F k in at_top. k \ x" for x::real by (simp add: eventually_ge_at_top) then have "\x\{a..}. ((\k. ?f k x) \ f x) at_top" by (auto intro!: Lim_transform_eventually[OF tendsto_const] simp: indicator_def eventually_at_top_linorder) ultimately have "((\k. integral {a..} (?f k)) \ integral {a ..} f) at_top" by (rule dominated_convergence_at_top) (auto simp: indicator_def) also have "(\k. integral {a..} (?f k)) = (\k. integral {a..k} f)" by (auto intro!: ext integral_indicator_eq_restriction f) finally show ?thesis . qed lemma norm_integrable_onI: "(\x. norm (f x)) integrable_on S" if "f absolutely_integrable_on S" for f::"'a::euclidean_space\'b::euclidean_space" using that by (auto simp: absolutely_integrable_on_def) lemma has_integral_improper_at_topI: fixes f::"real \ 'a::banach" assumes I: "\\<^sub>F k in at_top. (f has_integral I k) {a..k}" assumes J: "(I \ J) at_top" shows "(f has_integral J) {a..}" apply (subst has_integral') proof (auto, goal_cases) case (1 e) from tendstoD[OF J \0 < e\] have "\\<^sub>F x in at_top. dist (I x) J < e" . moreover have "\\<^sub>F x in at_top. (x::real) > 0" by simp moreover have "\\<^sub>F x in at_top. (x::real) > - a"\\TODO: this seems to be strange?\ by simp moreover note I ultimately have "\\<^sub>F x in at_top. x > 0 \ x > - a \ dist (I x) J < e \ (f has_integral I x) {a..x}" by eventually_elim auto then obtain k where k: "\b\k. norm (I b - J) < e" "k > 0" "k > - a" and I: "\c. c \ k \ (f has_integral I c) {a..c}" by (auto simp: eventually_at_top_linorder dist_norm) show ?case apply (rule exI[where x=k]) apply (auto simp: \0 < k\) subgoal premises prems for b c proof - have ball_eq: "ball 0 k = {-k <..< k}" by (auto simp: abs_real_def split: if_splits) from prems \0 < k\ have "c \ 0" "b \ 0" by (auto simp: subset_iff) with prems \0 < k\ have "c \ k" apply (auto simp: ball_eq) apply (auto simp: subset_iff) apply (drule spec[where x="(c + k)/2"]) - apply (auto simp: sign_simps not_less) + apply (auto simp: algebra_split_simps not_less) using \0 \ c\ by linarith then have "norm (I c - J) < e" using k by auto moreover from prems \0 < k\ \c \ 0\ \b \ 0\ \c \ k\ \k > - a\ have "a \ b" apply (auto simp: ball_eq) apply (auto simp: subset_iff) by (meson \b \ 0\ less_eq_real_def minus_less_iff not_le order_trans) have "((\x. if x \ cbox a c then f x else 0) has_integral I c) (cbox b c)" apply (subst has_integral_restrict_closed_subintervals_eq) using I[of c] prems \a \ b\ \k \ c\ by (auto ) from negligible_empty _ this have "((\x. if a \ x then f x else 0) has_integral I c) (cbox b c)" by (rule has_integral_spike) auto ultimately show ?thesis by (intro exI[where x="I c"]) auto qed done qed lemma has_integral_improperE: fixes f::"real \ 'a::euclidean_space" assumes I: "(f has_integral I) {a..}" assumes ai: "f absolutely_integrable_on {a..}" obtains J where "\k. (f has_integral J k) {a..k}" "(J \ I) at_top" proof - define J where "J k = integral {a .. k} f" for k have "(f has_integral J k) {a..k}" for k unfolding J_def by (force intro: integrable_on_subinterval has_integral_integrable[OF I]) moreover have I_def[symmetric]: "integral {a..} f = I" using I by auto from improper_integral_at_top[OF ai] have "(J \ I) at_top" unfolding J_def I_def . ultimately show ?thesis .. qed subsection \Miscellaneous\ lemma AE_BallI: "AE x\X in F. P x" if "\x \ X. P x" using that by (intro always_eventually) auto lemma bounded_le_Sup: assumes "bounded (f ` S)" shows "\x\S. norm (f x) \ Sup (norm ` f ` S)" by (auto intro!: cSup_upper bounded_imp_bdd_above simp: bounded_norm_comp assms) end \ No newline at end of file diff --git a/thys/LinearQuantifierElim/Thys/FRE.thy b/thys/LinearQuantifierElim/Thys/FRE.thy --- a/thys/LinearQuantifierElim/Thys/FRE.thy +++ b/thys/LinearQuantifierElim/Thys/FRE.thy @@ -1,205 +1,205 @@ (* Author: Tobias Nipkow, 2007 *) theory FRE imports LinArith begin subsection\Ferrante-Rackoff \label{sec:FRE}\ text\This section formalizes a slight variant of Ferrante and Rackoff's algorithm~\cite{FerranteR-SIAM75}. We consider equalities separately, which improves performance.\ fun between :: "real * real list \ real * real list \ real * real list" where "between (r,cs) (s,ds) = ((r+s)/2, (1/2) *\<^sub>s (cs+ds))" definition FR\<^sub>1 :: "atom fm \ atom fm" where "FR\<^sub>1 \ = (let as = R.atoms\<^sub>0 \; lbs = lbounds as; ubs = ubounds as; ebs = ebounds as; intrs = [subst \ (between l u) . l \ lbs, u \ ubs] in list_disj (inf\<^sub>- \ # inf\<^sub>+ \ # intrs @ map (subst \) ebs))" lemma dense_interval: assumes "finite L" "finite U" "l \ L" "u \ U" "l < x" "x < u" "P(x::real)" and dense: "\y l u. \ \y\{l<.. L; \y\{x<.. U; l \ P y" shows "\l\L.\u\U. l (\y. l y P y)" proof - let ?L = "{l:L. l < x}" let ?U = "{u:U. x < u}" let ?ll = "Max ?L" let ?uu = "Min ?U" have "?L \ {}" using \l \ L\ \l by (blast intro:order_less_imp_le) moreover have "?U \ {}" using \u:U\ \x by (blast intro:order_less_imp_le) ultimately have "\y. ?ll y y \ L" "\y. x y y \ U" using \finite L\ \finite U\ by force+ moreover have "?ll \ L" proof show "?ll \ ?L" using \finite L\ Max_in[OF _ \?L \ {}\] by simp show "?L \ L" by blast qed moreover have "?uu \ U" proof show "?uu \ ?U" using \finite U\ Min_in[OF _ \?U \ {}\] by simp show "?U \ U" by blast qed moreover have "?ll < x" using \finite L\ \?L \ {}\ by simp moreover have "x < ?uu" using \finite U\ \?U \ {}\ by simp moreover have "?ll < ?uu" using \?ll \x by arith ultimately show ?thesis using \l < x\ \x < u\ \?L \ {}\ \?U \ {}\ by(blast intro!:dense greaterThanLessThan_iff[THEN iffD1]) qed lemma dense: "\ nqfree f; \y\{l<.. LB f xs; \y\{x<.. UB f xs; l < x; x < u; x \ EQ f xs; R.I f (x#xs); l < y; y < u\ \ R.I f (y#xs)" proof(induct f) case (Atom a) show ?case proof (cases a) case (Less r cs) show ?thesis proof (cases cs) case Nil thus ?thesis using Atom Less by (simp add:depends\<^sub>R_def) next case (Cons c cs) hence "r < c*x + \cs,xs\" using Atom Less by simp { assume "c=0" hence ?thesis using Atom Less Cons by simp } moreover { assume "c<0" hence "x < (r - \cs,xs\)/c" (is "_ < ?u") using \r < c*x + \cs,xs\\ by (simp add: field_simps) have ?thesis proof (rule ccontr) assume "\ R.I (Atom a) (y#xs)" hence "?u \ y" using Atom Less Cons \c<0\ by (auto simp add: field_simps) hence "?u < u" using \y by simp with \x show False using Atom Less Cons \c<0\ by(auto simp:depends\<^sub>R_def) qed } moreover { assume "c>0" hence "x > (r - \cs,xs\)/c" (is "_ > ?l") using \r < c*x + \cs,xs\\ by (simp add: field_simps) have ?thesis proof (rule ccontr) assume "\ R.I (Atom a) (y#xs)" hence "?l \ y" using Atom Less Cons \c>0\ by (auto simp add: field_simps) hence "?l > l" using \y>l\ by simp with \?l show False using Atom Less Cons \c>0\ by (auto simp:depends\<^sub>R_def) qed } ultimately show ?thesis by force qed next case (Eq r cs) show ?thesis proof (cases cs) case Nil thus ?thesis using Atom Eq by (simp add:depends\<^sub>R_def) next case (Cons c cs) hence "r = c*x + \cs,xs\" using Atom Eq by simp { assume "c=0" hence ?thesis using Atom Eq Cons by simp } moreover { assume "c\0" hence ?thesis using \r = c*x + \cs,xs\\ Atom Eq Cons \l \y by(auto simp: ac_simps depends\<^sub>R_def split:if_splits) } ultimately show ?thesis by force qed qed next case (And f1 f2) thus ?case by auto (metis (no_types, hide_lams))+ next case (Or f1 f2) thus ?case by auto (metis (no_types, hide_lams))+ qed fastforce+ theorem I_FR\<^sub>1: assumes "nqfree \" shows "R.I (FR\<^sub>1 \) xs = (\x. R.I \ (x#xs))" (is "?FR = ?EX") proof assume ?FR { assume "R.I (inf\<^sub>- \) xs" hence ?EX using \?FR\ min_inf[OF \nqfree \\, where xs=xs] by(auto simp add:FR\<^sub>1_def) } moreover { assume "R.I (inf\<^sub>+ \) xs" hence ?EX using \?FR\ plus_inf[OF \nqfree \\, where xs=xs] by(auto simp add:FR\<^sub>1_def) } moreover { assume "\x \ EQ \ xs. R.I \ (x#xs)" hence ?EX using \?FR\ by(auto simp add:FR\<^sub>1_def) } moreover { assume "\R.I (inf\<^sub>- \) xs \ \R.I (inf\<^sub>+ \) xs \ (\x\EQ \ xs. \R.I \ (x#xs))" with \?FR\ obtain r cs s ds where "R.I (subst \ (between (r,cs) (s,ds))) xs" by(auto simp: FR\<^sub>1_def eval_def diff_divide_distrib set_ebounds I_subst \nqfree \\) blast hence "R.I \ (eval (between (r,cs) (s,ds)) xs # xs)" by(simp add:I_subst \nqfree \\) hence ?EX .. } ultimately show ?EX by blast next assume ?EX then obtain x where x: "R.I \ (x#xs)" .. { assume "R.I (inf\<^sub>- \) xs \ R.I (inf\<^sub>+ \) xs" hence ?FR by(auto simp:FR\<^sub>1_def) } moreover { assume "x \ EQ \ xs" then obtain r cs where "(r,cs) \ set(ebounds(R.atoms\<^sub>0 \)) \ x = r + \cs,xs\" by(force simp:set_ebounds field_simps) moreover hence "R.I (subst \ (r,cs)) xs" using x by(auto simp: I_subst \nqfree \\ eval_def) ultimately have ?FR by(force simp:FR\<^sub>1_def) } moreover { assume "\ R.I (inf\<^sub>- \) xs" and "\ R.I (inf\<^sub>+ \) xs" and "x \ EQ \ xs" obtain l where "l \ LB \ xs" "l < x" using LBex[OF \nqfree \\ x \\ R.I (inf\<^sub>- \) xs\ \x \ EQ \ xs\] .. obtain u where "u \ UB \ xs" "x < u" using UBex[OF \nqfree \\ x \\ R.I (inf\<^sub>+ \) xs\ \x \ EQ \ xs\] .. have "\l\LB \ xs. \u\UB \ xs. l (\y. l < y \ y < u \ R.I \ (y#xs))" using dense_interval[where P = "\x. R.I \ (x#xs)", OF finite_LB finite_UB \l:LB \ xs\ \u:UB \ xs\ \l \x x] x dense[OF \nqfree \\ _ _ _ _ \x \ EQ \ xs\] by simp then obtain r c cs s d ds where "Less r (c # cs) \ set (R.atoms\<^sub>0 \)" "Less s (d # ds) \ set (R.atoms\<^sub>0 \)" "\y. (r - \cs,xs\) / c < y \ y < (s - \ds,xs\) / d \ R.I \ (y # xs)" and *: "c > 0" "d < 0" "(r - \cs,xs\) / c < (s - \ds,xs\) / d" by blast moreover have "(r - \cs,xs\) / c < eval (between (r / c, (-1 / c) *\<^sub>s cs) (s / d, (-1 / d) *\<^sub>s ds)) xs" (is ?P) and "eval (between (r / c, (-1 / c) *\<^sub>s cs) (s / d, (-1 / d) *\<^sub>s ds)) xs < (s - \ds,xs\) / d" (is ?Q) proof - from * have [simp]: "c * (c * (d * (d * 4))) > 0" - by (simp add: sign_simps) + by (simp add: algebra_split_simps) from * have "c * s + d * \cs,xs\ < d * r + c * \ds,xs\" by (simp add: field_simps) with * have "(2 * c * c * d) * (d * r + c * \ds,xs\) < (2 * c * c * d) * (c * s + d * \cs,xs\)" and "(2 * c * d * d) * (c * s + d * \cs,xs\) < (2 * c * d * d) * (d * r + c * \ds,xs\)" by simp_all with * show ?P and ?Q by (auto simp add: field_simps eval_def iprod_left_add_distrib) qed ultimately have ?FR by (fastforce simp: FR\<^sub>1_def bex_Un set_lbounds set_ubounds set_ebounds I_subst \nqfree \\) } ultimately show ?FR by blast qed definition "FR = R.lift_nnf_qe FR\<^sub>1" lemma qfree_FR\<^sub>1: "nqfree \ \ qfree (FR\<^sub>1 \)" apply(simp add:FR\<^sub>1_def) apply(rule qfree_list_disj) apply(auto simp:qfree_min_inf qfree_plus_inf set_ubounds set_lbounds set_ebounds image_def qfree_map_fm) done theorem I_FR: "R.I (FR \) xs = R.I \ xs" by(simp add:I_FR\<^sub>1 FR_def R.I_lift_nnf_qe qfree_FR\<^sub>1) theorem qfree_FR: "qfree (FR \)" by(simp add:FR_def R.qfree_lift_nnf_qe qfree_FR\<^sub>1) end diff --git a/thys/Network_Security_Policy_Verification/Lib/ML_GraphViz.thy b/thys/Network_Security_Policy_Verification/Lib/ML_GraphViz.thy --- a/thys/Network_Security_Policy_Verification/Lib/ML_GraphViz.thy +++ b/thys/Network_Security_Policy_Verification/Lib/ML_GraphViz.thy @@ -1,183 +1,155 @@ theory ML_GraphViz -imports ML_GraphViz_Config +imports Main begin - -ML_val\ - val _ = writeln ("using `"^Graphviz_Platform_Config.executable_pdf_viewer^"` as pdf viewer and `"^ - Graphviz_Platform_Config.executable_dot^"` to render graphs."); -\ - - ML \ (*should we open a pdf viewer to display the generated graph?*) datatype open_viewer = DoNothing | OpenImmediately | AskTimeouted of real signature GRAPHVIZ = sig val open_viewer: open_viewer Unsynchronized.ref (*function to modify the printing of a node name*) val default_tune_node_format: term -> string -> string (* edges is a term of type ('a \ 'a) list *) (* @{context} default_tune_node_format (edges_format \ edges)list*) val visualize_graph: Proof.context -> (term -> string -> string) -> term -> unit (* @{context} default_tune_node_format (edges_format \ edges)list graphviz_header*) val visualize_graph_pretty: Proof.context -> (term -> string -> string) -> (string * term) list -> string-> unit (* helper function. @{context} tune_node_format node *) val node_to_string: Proof.context -> (term -> string -> string) -> term -> string val term_to_string: Proof.context -> term -> string; val term_to_string_safe: Proof.context -> term -> string; val term_to_string_html: Proof.context -> term -> string; end structure Graphviz: GRAPHVIZ = struct (*if set to `DoNothing`, graphviz will not be run and not pdf will be opened. Include ML_GraphViz_Disable.thy to run in batch mode.*) -val open_viewer = Unsynchronized.ref OpenImmediately +val open_viewer = (* FIXME avoid mutable state *) + Unsynchronized.ref (if getenv "ISABELLE_DOT" = "" then DoNothing else OpenImmediately) val default_tune_node_format: term -> string -> string = (fn _ => I) -fun write_to_tmpfile (t: string): Path.T = - let - val p = Isabelle_System.create_tmp_path "graphviz" "graph_tmp.dot" - in - writeln ("using tmpfile " ^ Path.print p); File.write p (t^"\n"); p - end - fun evaluate_term (ctxt: Proof.context) edges = case Code_Evaluation.dynamic_value ctxt edges of SOME x => x | NONE => error "ML_GraphViz: failed to evaluate term" fun is_valid_char c = (c <= #"z" andalso c >= #"a") orelse (c <= #"Z" andalso c >= #"A") orelse (c <= #"9" andalso c >= #"0") val sanitize_string = String.map (fn c => if is_valid_char c then c else #"_") fun term_to_string ctxt t = let val ctxt' = Config.put show_markup false ctxt; in Print_Mode.setmp [] (Syntax.string_of_term ctxt') t end; fun term_to_string_safe ctxt t = let val str = term_to_string ctxt t in if sanitize_string str <> str then (warning ("String "^str^" contains invalid characters!"); sanitize_string str) else str end; local val sanitize_string_html = String.map (fn c => if (is_valid_char c orelse c = #" " orelse (c <= #"/" andalso c >= #"(") orelse c = #"|" orelse c = #"=" orelse c = #"?" orelse c = #"!" orelse c = #"_" orelse c = #"[" orelse c = #"]") then c else #"_") in fun term_to_string_html ctxt (n: term) : string = let val str = term_to_string ctxt n in if sanitize_string_html str <> str then (warning ("String "^str^" contains invalid characters!"); sanitize_string_html str) else str end end; fun node_to_string ctxt (tune_node_format: term -> string -> string) (n: term) : string = n |> term_to_string ctxt |> tune_node_format n handle Subscript => error ("Subscript Exception in node_to_string for string " ^ (Pretty.string_of (Syntax.pretty_term ctxt n))); local - - (* viz is graphiz command, e.g. dot - viewer is a PDF viewer, e.g. xdg-open - retuns return code of bash command. - noticeable side effect: generated pdf file is not deleted (maybe still open in editor)*) - fun paint_graph (viewer: string) (viz: string) (f: Path.T) = - if (Isabelle_System.bash ("which "^viz)) <> 0 then - (*TODO: `which` on windows?*) - error "ML_GraphViz: Graphviz command not found" - else if (Isabelle_System.bash ("which "^viewer)) <> 0 then - error "ML_GraphViz: viewer command not found" + fun display_graph graph = + if getenv "ISABELLE_DOT" = "" then + error "Missing $ISABELLE_DOT settings variable (Graphviz \"dot\" executable)" else - let - val base = Path.base f; - val base_pdf = base |> Path.ext "pdf"; - (*First cd to the temp directory, then only call the commands with relative paths. - This is a Windows workaround if the Windows (not cygwin) version of graphviz is installed: - It does not understand paths such as /tmp/isabelle/.., it wants C:\tmp\.. - Hence, we cd to the tmp directory and only use relative filenames henceforth.*) - val cmd = - "cd " ^ File.bash_path (Path.dir f) ^ "; " ^ - viz ^ " -o "^ File.bash_path base_pdf ^ " -Tpdf " ^ File.bash_path base ^ - " && " ^ viewer ^ " " ^ File.bash_path base_pdf; - in - writeln ("executing: "^cmd); - Isabelle_System.bash cmd; - File.rm f (*cleanup dot file, PDF file will still exist*) - (*some pdf viewers do not like it if we delete the pdf file they are currently displaying*) - end + Isabelle_System.with_tmp_file "graphviz" "dot" (fn graph_file => + let + val _ = File.write graph_file graph; + val pdf_file = Path.explode "$ISABELLE_HOME_USER/graphviz.pdf"; + val _ = + (Isabelle_System.bash o cat_lines) + ["set -e", + "cd " ^ File.bash_path (Path.dir graph_file), + "\"$ISABELLE_DOT\" -o " ^ Bash_Syntax.string (File.platform_path pdf_file) ^ + " -Tpdf " ^ Bash_Syntax.string (File.platform_path graph_file), + "\"$PDF_VIEWER\" " ^ File.bash_path pdf_file ^ " &"]; + in () end); fun format_dot_edges ctxt tune_node_format trm = let fun format_node t = let val str = node_to_string ctxt tune_node_format t in if sanitize_string str <> str then (warning ("Node "^str^" contains invalid characters!"); sanitize_string str) else str end; fun format_dot_edge (t1, t2) = format_node t1 ^ " -> " ^ format_node t2 ^ ";\n" in map format_dot_edge trm end fun apply_dot_header header edges = "digraph graphname {\n#header\n" ^ header ^"\n#edges\n\n"^ implode edges ^ "}" in fun visualize_graph_pretty ctxt tune_node_format Es (header: string) = let val evaluated_edges = map (fn (str, t) => (str, evaluate_term ctxt t)) Es; val edge_to_string = HOLogic.dest_list #> map HOLogic.dest_prod #> format_dot_edges ctxt tune_node_format #> implode; val formatted_edges = map (fn (str, t) => str ^ "\n" ^ edge_to_string t) evaluated_edges; - val execute_command = fn _ => apply_dot_header header formatted_edges - |> write_to_tmpfile - |> paint_graph Graphviz_Platform_Config.executable_pdf_viewer Graphviz_Platform_Config.executable_dot; + fun execute_command () = display_graph (apply_dot_header header formatted_edges); in case !open_viewer of DoNothing => writeln "visualization disabled (Graphviz.open_viewer)" | OpenImmediately => execute_command () | AskTimeouted wait_seconds => let val (text, promise) = Active.dialog_text (); val _ = writeln ("Run Grpahviz and display pdf? " ^ text "yes" ^ "/" ^ text "no" ^ " (execution paused)") in Timeout.apply (seconds wait_seconds) (fn _ => let val m = (Future.join promise) in (if m = "yes" then execute_command () else (writeln "no")) end ) () end handle Timeout.TIMEOUT _ => (writeln ("Timeouted. You did not klick yes/no. Killed to continue. " ^ "If you want to see the pdf, just re-run this and klick yes.")) end end fun visualize_graph ctxt tune_node_format edges = visualize_graph_pretty ctxt tune_node_format [("", edges)] "#TODO add header here" end; \ end diff --git a/thys/Network_Security_Policy_Verification/Lib/ML_GraphViz_Config.thy b/thys/Network_Security_Policy_Verification/Lib/ML_GraphViz_Config.thy deleted file mode 100644 --- a/thys/Network_Security_Policy_Verification/Lib/ML_GraphViz_Config.thy +++ /dev/null @@ -1,38 +0,0 @@ -theory ML_GraphViz_Config -imports Main -begin - -ML\ - -signature GRAPHVIZ_PLATFORM_CONFIG = -sig - val executable_dot: string; - val executable_pdf_viewer: string; -end - -structure Graphviz_Platform_Config: GRAPHVIZ_PLATFORM_CONFIG = -struct - (*Change your system config here*) - val (executable_dot: string, executable_pdf_viewer: string) = ( - case getenv "ISABELLE_PLATFORM_FAMILY" of - "linux" => ("dot", getenv "PDF_VIEWER") (*tested on ubuntu 14.04, graphviz 2.36*) - | "macos" => ("dot", getenv "PDF_VIEWER") - | "windows" => ("dot", getenv "PDF_VIEWER") (*tested with graphviz-2.38.msi, works, needed to add the bin directory to $PATH manually*) - | _ => raise Fail "$ISABELLE_PLATFORM_FAMILY: cannot determine operating system" - ); - - local - fun check_executable e = - if Isabelle_System.bash ("which "^e) = 0 then - () (* `which` already printed the path *) - else - warning ("Command not available or not in $PATH: "^e); - in - val _ = check_executable executable_pdf_viewer; - val _ = check_executable executable_dot; - end - -end -\ - -end diff --git a/thys/Ordinary_Differential_Equations/IVP/Cones.thy b/thys/Ordinary_Differential_Equations/IVP/Cones.thy --- a/thys/Ordinary_Differential_Equations/IVP/Cones.thy +++ b/thys/Ordinary_Differential_Equations/IVP/Cones.thy @@ -1,694 +1,694 @@ theory Cones imports "HOL-Analysis.Analysis" Triangle.Triangle "../ODE_Auxiliarities" begin lemma arcsin_eq_zero_iff[simp]: "-1 \ x \ x \ 1 \ arcsin x = 0 \ x = 0" using sin_arcsin by fastforce definition conemem :: "'a::real_vector \ 'a \ real \ 'a" where "conemem u v t = cos t *\<^sub>R u + sin t *\<^sub>R v" definition "conesegment u v = conemem u v ` {0.. pi / 2}" lemma bounded_linear_image_conemem: assumes "bounded_linear F" shows "F (conemem u v t) = conemem (F u) (F v) t" proof - from assms interpret bounded_linear F . show ?thesis by (auto simp: conemem_def[abs_def] cone_hull_expl closed_segment_def add scaleR) qed lemma bounded_linear_image_conesegment: assumes "bounded_linear F" shows "F ` conesegment u v = conesegment (F u) (F v)" proof - from assms interpret bounded_linear F . show ?thesis apply (auto simp: conesegment_def conemem_def[abs_def] cone_hull_expl closed_segment_def add scaleR) apply (auto simp: add[symmetric] scaleR[symmetric]) done qed (* This is vangle in $AFP/Triangles/Angles *) lemma discriminant: "a * x\<^sup>2 + b * x + c = (0::real) \ 0 \ b\<^sup>2 - 4 * a * c" by (sos "(((A<0 * R<1) + (R<1 * (R<1 * [2*a*x + b]^2))))") lemma quadratic_eq_factoring: assumes D: "D = b\<^sup>2 - 4 * a * c" assumes nn: "0 \ D" assumes x1: "x\<^sub>1 = (-b + sqrt D) / (2 * a)" assumes x2: "x\<^sub>2 = (-b - sqrt D) / (2 * a)" assumes a: "a \ 0" shows "a * x\<^sup>2 + b * x + c = a * (x - x\<^sub>1) * (x - x\<^sub>2)" using nn by (simp add: D x1 x2) (simp add: assms algebra_simps power2_eq_square power3_eq_cube divide_simps) lemma quadratic_eq_zeroes_iff: assumes D: "D = b\<^sup>2 - 4 * a * c" assumes x1: "x\<^sub>1 = (-b + sqrt D) / (2 * a)" assumes x2: "x\<^sub>2 = (-b - sqrt D) / (2 * a)" assumes a: "a \ 0" shows "a * x\<^sup>2 + b * x + c = 0 \ (D \ 0 \ (x = x\<^sub>1 \ x = x\<^sub>2))" (is "?z \ _") using quadratic_eq_factoring[OF D _ x1 x2 a, of x] discriminant[of a x b c] a by (auto simp: D) lemma quadratic_ex_zero_iff: "(\x. a * x\<^sup>2 + b * x + c = 0) \ (a \ 0 \ b\<^sup>2 - 4 * a * c \ 0 \ a = 0 \ (b = 0 \ c = 0))" for a b c::real apply (cases "a = 0") subgoal by (auto simp: intro: exI[where x="- c / b"]) subgoal by (subst quadratic_eq_zeroes_iff[OF refl refl refl]) auto done lemma Cauchy_Schwarz_eq_iff: shows "(inner x y)\<^sup>2 = inner x x * inner y y \ ((\k. x = k *\<^sub>R y) \ y = 0)" proof safe assume eq: "(x \ y)\<^sup>2 = x \ x * (y \ y)" and "y \ 0" define f where "f \ \l. inner (x - l *\<^sub>R y) (x - l *\<^sub>R y)" have f_quadratic: "f l = inner y y * l\<^sup>2 + - 2 * inner x y * l + inner x x" for l by (auto simp: f_def algebra_simps power2_eq_square inner_commute) have "\l. f l = 0" unfolding f_quadratic quadratic_ex_zero_iff using \y \ 0\ by (auto simp: eq) then show "(\k. x = k *\<^sub>R y)" by (auto simp: f_def) qed (auto simp: power2_eq_square) lemma Cauchy_Schwarz_strict_ineq: "(inner x y)\<^sup>2 < inner x x * inner y y" if "y \ 0" "\k. x \ k *\<^sub>R y" apply (rule neq_le_trans) subgoal using that unfolding Cauchy_Schwarz_eq_iff by auto subgoal by (rule Cauchy_Schwarz_ineq) done lemma Cauchy_Schwarz_eq2_iff: "\inner x y\ = norm x * norm y \ ((\k. x = k *\<^sub>R y) \ y = 0)" using Cauchy_Schwarz_eq_iff[of x y] by (subst power_eq_iff_eq_base[symmetric, where n = 2]) (simp_all add: dot_square_norm power_mult_distrib) lemma Cauchy_Schwarz_strict_ineq2: "\inner x y\ < norm x * norm y" if "y \ 0" "\k. x \ k *\<^sub>R y" apply (rule neq_le_trans) subgoal using that unfolding Cauchy_Schwarz_eq2_iff by auto subgoal by (rule Cauchy_Schwarz_ineq2) done lemma gt_minus_one_absI: "abs k < 1 \ - 1 < k" for k::real by auto lemma gt_one_absI: "abs k < 1 \ k < 1" for k::real by auto lemma abs_impossible: "\y1\ < x1 \ \y2\ < x2 \ x1 * x2 + y1 * y2 \ 0" for x1 x2::real proof goal_cases case 1 have "- y1 * y2 \ abs y1 * abs y2" by (metis abs_ge_minus_self abs_mult mult.commute mult_minus_right) also have "\ < x1 * x2" apply (rule mult_strict_mono) using 1 by auto finally show ?case by auto qed lemma vangle_eq_arctan_minus:\ \TODO: generalize?!\ assumes ij: "i \ Basis" "j \ Basis" and ij_neq: "i \ j" assumes xy1: "\y1\ < x1" assumes xy2: "\y2\ < x2" assumes less: "y2 / x2 > y1 / x1" shows "vangle (x1 *\<^sub>R i + y1 *\<^sub>R j) (x2 *\<^sub>R i + y2 *\<^sub>R j) = arctan (y2 / x2) - arctan (y1 / x1)" (is "vangle ?u ?v = _") proof - from assms have less2: "x2 * y1 - x1 * y2 < 0" by (auto simp: divide_simps abs_real_def algebra_simps split: if_splits) have norm_eucl: "norm (x *\<^sub>R i + y *\<^sub>R j) = sqrt ((norm x)\<^sup>2 + (norm y)\<^sup>2)" for x y apply (subst norm_eq_sqrt_inner) using ij ij_neq by (auto simp: inner_simps inner_Basis power2_eq_square) have nonzeroes: "x1 *\<^sub>R i + y1 *\<^sub>R j \ 0" "x2 *\<^sub>R i + y2 *\<^sub>R j \ 0" apply (auto simp: euclidean_eq_iff[where 'a='a] inner_simps intro!: bexI[where x=i]) using assms by (auto simp: inner_Basis) have indep: "x1 *\<^sub>R i + y1 *\<^sub>R j \ k *\<^sub>R (x2 *\<^sub>R i + y2 *\<^sub>R j)" for k proof assume "x1 *\<^sub>R i + y1 *\<^sub>R j = k *\<^sub>R (x2 *\<^sub>R i + y2 *\<^sub>R j)" then have "x1 / x2 = k" "y1 = k * y2" using ij ij_neq xy1 xy2 apply (auto simp: abs_real_def divide_simps algebra_simps euclidean_eq_iff[where 'a='a] inner_simps split: if_splits) by (auto simp: inner_Basis split: if_splits) then have "y1 = x1 / x2 * y2" by simp with less show False using xy1 by (auto split: if_splits) qed have "((x1\<^sup>2 + y1\<^sup>2) * (x2\<^sup>2 + y2\<^sup>2) * (1 - ((x1 *\<^sub>R i + y1 *\<^sub>R j) \ (x2 *\<^sub>R i + y2 *\<^sub>R j))\<^sup>2 / ((x1\<^sup>2 + y1\<^sup>2) * (x2\<^sup>2 + y2\<^sup>2)))) = ((x1\<^sup>2 + y1\<^sup>2) * (x2\<^sup>2 + y2\<^sup>2) * (1 - (x1 * x2 + y1 * y2)\<^sup>2 / ((x1\<^sup>2 + y1\<^sup>2) * (x2\<^sup>2 + y2\<^sup>2))))" using ij_neq ij by (auto simp: algebra_simps divide_simps inner_simps inner_Basis) also have "\ = (x1\<^sup>2 + y1\<^sup>2) * (x2\<^sup>2 + y2\<^sup>2) - (x1 * x2 + y1 * y2)\<^sup>2" unfolding right_diff_distrib by simp also have "\ = (x2 * y1 - x1 * y2)^2" by (auto simp: algebra_simps power2_eq_square) also have "sqrt \ = \x2 * y1 - x1 * y2\" by simp also have "\ = x1 * y2 - x2 * y1" using less2 by (simp add: abs_real_def) finally have sqrt_eq: "sqrt ((x1\<^sup>2 + y1\<^sup>2) * (x2\<^sup>2 + y2\<^sup>2) * (1 - ((x1 *\<^sub>R i + y1 *\<^sub>R j) \ (x2 *\<^sub>R i + y2 *\<^sub>R j))\<^sup>2 / ((x1\<^sup>2 + y1\<^sup>2) * (x2\<^sup>2 + y2\<^sup>2)))) = x1 * y2 - x2 * y1" . show ?thesis using ij xy1 xy2 unfolding vangle_def apply (subst arccos_arctan) subgoal apply (rule gt_minus_one_absI) apply (simp add: ) apply (subst pos_divide_less_eq) subgoal apply (rule mult_pos_pos) using nonzeroes by auto subgoal apply simp apply (rule Cauchy_Schwarz_strict_ineq2) using nonzeroes indep by auto done subgoal apply (rule gt_one_absI) apply (simp add: ) apply (subst pos_divide_less_eq) subgoal apply (rule mult_pos_pos) using nonzeroes by auto subgoal apply simp apply (rule Cauchy_Schwarz_strict_ineq2) using nonzeroes indep by auto done subgoal apply (auto simp: nonzeroes) apply (subst (3) diff_conv_add_uminus) apply (subst arctan_minus[symmetric]) apply (subst arctan_add) apply force apply force apply (subst arctan_inverse[symmetric]) subgoal apply (rule divide_pos_pos) subgoal apply (auto simp add: inner_simps inner_Basis algebra_simps ) apply (thin_tac "_ \ Basis")+ apply (thin_tac "j = i") apply (sos "((((A<0 * (A<1 * (A<2 * A<3))) * R<1) + ((A<=0 * (A<0 * (A<2 * R<1))) * (R<1 * [1]^2))))") apply (thin_tac "_ \ Basis")+ apply (thin_tac "j \ i") by (sos "((((A<0 * (A<1 * (A<2 * A<3))) * R<1) + (((A<2 * (A<3 * R<1)) * (R<1/3 * [y1]^2)) + (((A<1 * (A<3 * R<1)) * ((R<1/12 * [x2 + y1]^2) + (R<1/12 * [x1 + y2]^2))) + (((A<1 * (A<2 * R<1)) * (R<1/12 * [~1*x1 + x2 + y1 + y2]^2)) + (((A<0 * (A<3 * R<1)) * (R<1/12 * [~1*x1 + x2 + ~1*y1 + ~1*y2]^2)) + (((A<0 * (A<2 * R<1)) * ((R<1/12 * [x2 + ~1*y1]^2) + (R<1/12 * [~1*x1 + y2]^2))) + (((A<0 * (A<1 * R<1)) * (R<1/3 * [y2]^2)) + ((A<=0 * R<1) * (R<1/3 * [x1 + x2]^2))))))))))") subgoal apply (intro mult_pos_pos) using nonzeroes indep apply auto apply (rule gt_one_absI) apply (simp add: power_divide power_mult_distrib power2_norm_eq_inner) apply (rule Cauchy_Schwarz_strict_ineq) apply auto done done subgoal apply (rule arg_cong[where f=arctan]) using nonzeroes ij_neq apply (auto simp: norm_eucl) apply (subst real_sqrt_mult[symmetric]) apply (subst real_sqrt_mult[symmetric]) apply (subst real_sqrt_mult[symmetric]) apply (subst power_divide) apply (subst real_sqrt_pow2) apply simp apply (subst nonzero_divide_eq_eq) subgoal apply (auto simp: algebra_simps inner_simps inner_Basis) by (auto simp: algebra_simps divide_simps abs_real_def abs_impossible) apply (subst sqrt_eq) apply (auto simp: algebra_simps inner_simps inner_Basis) apply (auto simp: algebra_simps divide_simps abs_real_def abs_impossible) by (auto split: if_splits) done done qed lemma vangle_le_pi2: "0 \ u \ v \ vangle u v \ pi/2" unfolding vangle_def atLeastAtMost_iff apply (simp del: le_divide_eq_numeral1) apply (intro impI arccos_le_pi2 arccos_lbound) using Cauchy_Schwarz_ineq2[of u v] by (auto simp: algebra_simps) lemma inner_eq_vangle: "u \ v = cos (vangle u v) * (norm u * norm v)" by (simp add: cos_vangle) lemma vangle_scaleR_self: "vangle (k *\<^sub>R v) v = (if k = 0 \ v = 0 then pi / 2 else if k > 0 then 0 else pi)" "vangle v (k *\<^sub>R v) = (if k = 0 \ v = 0 then pi / 2 else if k > 0 then 0 else pi)" by (auto simp: vangle_def dot_square_norm power2_eq_square) lemma vangle_scaleR: "vangle (k *\<^sub>R v) w = vangle v w" "vangle w (k *\<^sub>R v) = vangle w v" if "k > 0" using that by (auto simp: vangle_def) lemma cos_vangle_eq_zero_iff_vangle: "cos (vangle u v) = 0 \ (u = 0 \ v = 0 \ u \ v = 0)" using Cauchy_Schwarz_ineq2[of u v] - by (auto simp: vangle_def divide_simps sign_simps split: if_splits) + by (auto simp: vangle_def divide_simps algebra_split_simps split: if_splits) lemma ortho_imp_angle_pi_half: "u \ v = 0 \ vangle u v = pi / 2" using orthogonal_iff_vangle[of u v] by (auto simp: orthogonal_def) lemma arccos_eq_zero_iff: "arccos x = 0 \ x = 1" if "-1 \ x" "x \ 1" using that apply auto using cos_arccos by fastforce lemma vangle_eq_zeroD: "vangle u v = 0 \ (\k. v = k *\<^sub>R u)" apply (auto simp: vangle_def split: if_splits) apply (subst (asm) arccos_eq_zero_iff) apply (auto simp: divide_simps mult_less_0_iff split: if_splits) apply (metis Real_Vector_Spaces.norm_minus_cancel inner_minus_left minus_le_iff norm_cauchy_schwarz) apply (metis norm_cauchy_schwarz) by (metis Cauchy_Schwarz_eq2_iff abs_of_pos inner_commute mult.commute mult_sign_intros(5) zero_less_norm_iff) lemma less_one_multI:\ \TODO: also in AA!\ 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 conemem_expansion_estimate: fixes u v u' v'::"'a::euclidean_space" assumes "t \ {0 .. pi / 2}" assumes angle_pos: "0 < vangle u v" "vangle u v < pi / 2" assumes angle_le: "(vangle u' v') \ (vangle u v)" assumes "norm u = 1" "norm v = 1" shows "norm (conemem u' v' t) \ min (norm u') (norm v') * norm (conemem u v t)" proof - define e_pre where "e_pre = min (norm u') (norm v')" let ?w = "conemem u v" let ?w' = "conemem u' v'" have cos_angle_le: "cos (vangle u' v') \ cos (vangle u v)" using angle_pos vangle_bounds by (auto intro!: cos_monotone_0_pi_le angle_le) have e_pre_le: "e_pre\<^sup>2 \ norm u' * norm v'" by (auto simp: e_pre_def min_def power2_eq_square intro: mult_left_mono mult_right_mono) have lt: "0 < 1 + 2 * (u \ v) * sin t * cos t" proof - have "\u \ v\ < norm u * norm v" apply (rule Cauchy_Schwarz_strict_ineq2) using assms apply auto apply (subst (asm) vangle_scaleR_self)+ by (auto simp: split: if_splits) then have "abs (u \ v * sin (2 * t)) < 1" using assms apply (auto simp add: abs_mult) apply (subst mult.commute) apply (rule less_one_multI) apply (auto simp add: abs_mult inner_eq_vangle ) by (auto simp: cos_vangle_eq_zero_iff_vangle dest!: ortho_imp_angle_pi_half) then show ?thesis by (subst mult.assoc sin_times_cos)+ auto qed have le: "0 \ 1 + 2 * (u \ v) * sin t * cos t" proof - have "\u \ v\ \ norm u * norm v" by (rule Cauchy_Schwarz_ineq2) then have "abs (u \ v * sin (2 * t)) \ 1" by (auto simp add: abs_mult assms intro!: mult_le_one) then show ?thesis by (subst mult.assoc sin_times_cos)+ auto qed have "(norm (?w t))\<^sup>2 = (cos t)\<^sup>2 *\<^sub>R (norm u)\<^sup>2 + (sin t)\<^sup>2 *\<^sub>R (norm v)\<^sup>2 + 2 * (u \ v) * sin t * cos t" by (auto simp: conemem_def algebra_simps power2_norm_eq_inner) (auto simp: power2_eq_square inner_commute) also have "\ = 1 + 2 * (u \ v) * sin t * cos t" by (auto simp: sin_squared_eq algebra_simps assms) finally have "(norm (conemem u v t))\<^sup>2 = 1 + 2 * (u \ v) * sin t * cos t" by simp moreover have "(norm (?w' t))\<^sup>2 = (cos t)\<^sup>2 *\<^sub>R (norm u')\<^sup>2 + (sin t)\<^sup>2 *\<^sub>R (norm v')\<^sup>2 + 2 * (u' \ v') * sin t * cos t" by (auto simp: conemem_def algebra_simps power2_norm_eq_inner) (auto simp: power2_eq_square inner_commute) ultimately have "(norm (?w' t) / norm (?w t))\<^sup>2 = ((cos t)\<^sup>2 *\<^sub>R (norm u')\<^sup>2 + (sin t)\<^sup>2 *\<^sub>R (norm v')\<^sup>2 + 2 * (u' \ v') * sin t * cos t) / (1 + 2 * (u \ v) * sin t * cos t)" (is "_ = (?a + ?b) / ?c") by (auto simp: divide_inverse power_mult_distrib) (auto simp: inverse_eq_divide power2_eq_square) also have "\ \ (e_pre\<^sup>2 + ?b) / ?c" apply (rule divide_right_mono) apply (rule add_right_mono) subgoal using assms e_pre_def apply (auto simp: min_def) subgoal by (auto simp: algebra_simps cos_squared_eq intro!: mult_right_mono power_mono) subgoal by (auto simp: algebra_simps sin_squared_eq intro!: mult_right_mono power_mono) done subgoal by (rule le) done also (xtrans) have inner_nonneg: "u' \ v' \ 0" using angle_le(1) angle_pos vangle_bounds[of u' v'] by (auto simp: inner_eq_vangle intro!: mult_nonneg_nonneg cos_ge_zero) from vangle_bounds[of u' v'] vangle_le_pi2[OF this] have u'v'e_pre: "u' \ v' \ cos (vangle u' v') * e_pre\<^sup>2" apply (subst inner_eq_vangle) apply (rule mult_left_mono) apply (rule e_pre_le) apply (rule cos_ge_zero) by auto have "(e_pre\<^sup>2 + ?b) / ?c \ (e_pre\<^sup>2 + 2 * (cos (vangle u' v') * e_pre\<^sup>2) * sin t * cos t) / ?c" (is "_ \ ?ddd") apply (intro divide_right_mono add_left_mono mult_right_mono mult_left_mono u'v'e_pre) using \t \ _\ by (auto intro!: mult_right_mono sin_ge_zero divide_right_mono le cos_ge_zero simp: sin_times_cos u'v'e_pre) also (xtrans) have "?ddd = e_pre\<^sup>2 * ((1 + 2 * cos (vangle u' v') * sin t * cos t) / ?c)" (is "_ = ?ddd") by (auto simp add: divide_simps algebra_simps) also (xtrans) have sc_ge_0: "0 \ sin t * cos t" using \t \ _\ by (auto simp: assms cos_angle_le intro!: mult_nonneg_nonneg sin_ge_zero cos_ge_zero) have "?ddd \ e_pre\<^sup>2" apply (subst mult_le_cancel_left1) apply (auto simp add: divide_simps split: if_splits) apply (rule mult_right_mono) using lt by (auto simp: assms inner_eq_vangle intro!: mult_right_mono sc_ge_0 cos_angle_le) finally (xtrans) have "(norm (conemem u' v' t))\<^sup>2 \ (e_pre * norm (conemem u v t))\<^sup>2" by (simp add: divide_simps power_mult_distrib split: if_splits) then show "norm (conemem u' v' t) \ e_pre * norm (conemem u v t)" using norm_imp_pos_and_ge power2_le_imp_le by blast qed lemma conemem_commute: "conemem a b t = conemem b a (pi / 2 - t)" if "0 \ t" "t \ pi / 2" using that by (auto simp: conemem_def cos_sin_eq algebra_simps) lemma conesegment_commute: "conesegment a b = conesegment b a" apply (auto simp: conesegment_def ) apply (subst conemem_commute) apply auto apply (subst conemem_commute) apply auto done definition "conefield u v = cone hull (conesegment u v)" lemma conefield_alt_def: "conefield u v = cone hull {u--v}" apply (auto simp: conesegment_def conefield_def cone_hull_expl in_segment) subgoal premises prems for c t proof - from prems have sc_pos: "sin t + cos t > 0" apply (cases "t = 0") subgoal by (rule add_nonneg_pos) auto subgoal by (auto intro!: add_pos_nonneg sin_gt_zero cos_ge_zero) done then have 1: "(sin t / (sin t + cos t) + cos t / (sin t + cos t)) = 1" by (auto simp: divide_simps) have "\c x. c > 0 \ 0 \ x \ x \ 1 \ c *\<^sub>R conemem u v t = (1 - x) *\<^sub>R u + x *\<^sub>R v" apply (auto simp: algebra_simps conemem_def) apply (rule exI[where x="1 / (sin t + cos t)"]) using prems by (auto intro!: exI[where x="(1 / (sin t + cos t) * sin t)"] sc_pos divide_nonneg_nonneg sin_ge_zero add_nonneg_nonneg cos_ge_zero simp: scaleR_add_left[symmetric] 1 divide_le_eq_1) then obtain d x where dx: "d > 0" "conemem u v t = (1 / d) *\<^sub>R ((1 - x) *\<^sub>R u + x *\<^sub>R v)" "0 \ x" "x \ 1" by (auto simp: eq_vector_fraction_iff) show ?thesis apply (rule exI[where x="c / d"]) using dx by (auto simp: intro!: divide_nonneg_nonneg prems ) qed subgoal premises prems for c t proof - let ?x = "arctan (t / (1 - t))" let ?s = "t / sin ?x" have *: "c *\<^sub>R ((1 - t) *\<^sub>R u + t *\<^sub>R v) = (c * ?s) *\<^sub>R (cos ?x *\<^sub>R u + sin ?x *\<^sub>R v)" if "0 < t" "t < 1" using that by (auto simp: scaleR_add_right sin_arctan cos_arctan divide_simps) show ?thesis apply (cases "t = 0") subgoal apply simp apply (rule exI[where x=c]) apply (rule exI[where x=u]) using prems by (auto simp: conemem_def[abs_def] intro!: image_eqI[where x=0]) subgoal apply (cases "t = 1") subgoal apply simp apply (rule exI[where x=c]) apply (rule exI[where x=v]) using prems by (auto simp: conemem_def[abs_def] intro!: image_eqI[where x="pi/2"]) subgoal apply (rule exI[where x="(c * ?s)"]) apply (rule exI[where x="(cos ?x *\<^sub>R u + sin ?x *\<^sub>R v)"]) using prems * arctan_ubound[of "t / (1 - t)"] apply (auto simp: conemem_def[abs_def] intro!: imageI) by (auto simp: scaleR_add_right sin_arctan) done done qed done lemma bounded_linear_image_cone_hull: assumes "bounded_linear F" shows "F ` (cone hull T) = cone hull (F ` T)" proof - from assms interpret bounded_linear F . show ?thesis apply (auto simp: conefield_def cone_hull_expl closed_segment_def add scaleR) apply (auto simp: ) apply (auto simp: add[symmetric] scaleR[symmetric]) done qed lemma bounded_linear_image_conefield: assumes "bounded_linear F" shows "F ` conefield u v = conefield (F u) (F v)" unfolding conefield_def using assms by (auto simp: bounded_linear_image_conesegment bounded_linear_image_cone_hull) lemma conefield_commute: "conefield x y = conefield y x" by (auto simp: conefield_def conesegment_commute) lemma convex_conefield: "convex (conefield x y)" by (auto simp: conefield_alt_def convex_cone_hull) lemma conefield_scaleRI: "v \ conefield (r *\<^sub>R x) y" if "v \ conefield x y" "r > 0" using that using \r > 0\ unfolding conefield_alt_def cone_hull_expl apply (auto simp: in_segment) proof goal_cases case (1 c u) let ?d = "c * (1 - u) / r + c * u" let ?t = "c * u / ?d" have "c * (1 - u) = ?d * (1 - ?t) * r" if "0 < u" using \0 < r\ that(1) 1(3,5) mult_pos_pos by (force simp: divide_simps ac_simps ring_distribs[symmetric]) then have eq1: "(c * (1 - u)) *\<^sub>R x = (?d * (1 - ?t) * r) *\<^sub>R x" if "0 < u" using that by simp have "c * u = ?d * ?t" if "u < 1" using \0 < r\ that(1) 1(3,4,5) mult_pos_pos apply (auto simp: divide_simps ac_simps ring_distribs[symmetric]) proof - assume "0 \ u" "0 < r" "1 - u + r * u = 0" "u < 1" then have False by (sos "((((A<0 * A<1) * R<1) + (([~1*r] * A=0) + ((A<=0 * R<1) * (R<1 * [r]^2)))))") then show "u = 0" by metis qed then have eq2: "(c * u) *\<^sub>R y = (?d * ?t) *\<^sub>R y" if "u < 1" using that by simp have *: "c *\<^sub>R ((1 - u) *\<^sub>R x + u *\<^sub>R y) = ?d *\<^sub>R ((1 - ?t) *\<^sub>R r *\<^sub>R x + ?t *\<^sub>R y)" if "0 < u" "u < 1" using that eq1 eq2 by (auto simp: algebra_simps) show ?case apply (cases "u = 0") subgoal using 1 by (intro exI[where x="c / r"] exI[where x="r *\<^sub>R x"]) auto apply (cases "u = 1") subgoal using 1 by (intro exI[where x="c"] exI[where x="y"]) (auto intro!: exI[where x=1]) subgoal apply (rule exI[where x="?d"]) apply (rule exI[where x="((1 - ?t) *\<^sub>R r *\<^sub>R x + ?t *\<^sub>R y)"]) apply (subst *) using 1 apply (auto intro!: exI[where x = ?t]) apply (auto simp: algebra_simps divide_simps) defer proof - assume a1: "c + c * (r * u) < c * u" assume a2: "0 \ c" assume a3: "0 \ u" assume a4: "u \ 0" assume a5: "0 < r" have "c + c * (r * u) \ c * u" using a1 less_eq_real_def by blast then show "c \ c * u" using a5 a4 a3 a2 by (metis (no_types) less_add_same_cancel1 less_eq_real_def mult_pos_pos order_trans real_scaleR_def real_vector.scale_zero_left) next assume a1: "0 \ c" assume a2: "u \ 1" have f3: "\x0. ((x0::real) < 1) = (\ 1 \ x0)" by auto have f4: "\x0. ((1::real) < x0) = (\ x0 \ 1)" by fastforce have "\x0 x1. ((x1::real) < x1 * x0) = (\ 0 \ x1 + - 1 * (x1 * x0))" by auto then have "(\r ra. ((r::real) < r * ra) = ((0 \ r \ 1 < ra) \ (r \ 0 \ ra < 1))) = (\r ra. (\ (0::real) \ r + - 1 * (r * ra)) = ((\ 0 \ r \ \ ra \ 1) \ (\ r \ 0 \ \ 1 \ ra)))" using f4 f3 by presburger then have "0 \ c + - 1 * (c * u)" using a2 a1 mult_less_cancel_left1 by blast then show "c * u \ c" by auto qed done qed lemma conefield_scaleRD: "v \ conefield x y" if "v \ conefield (r *\<^sub>R x) y" "r > 0" using conefield_scaleRI[OF that(1) positive_imp_inverse_positive[OF that(2)]] that(2) by auto lemma conefield_scaleR: "conefield (r *\<^sub>R x) y = conefield x y" if "r > 0" using conefield_scaleRD conefield_scaleRI that by blast lemma conefield_expansion_estimate: fixes u v::"'a::euclidean_space" and F::"'a \ 'a" assumes "t \ {0 .. pi / 2}" assumes angle_pos: "0 < vangle u v" "vangle u v < pi / 2" assumes angle_le: "vangle (F u) (F v) \ vangle u v" assumes "bounded_linear F" assumes "x \ conefield u v" shows "norm (F x) \ min (norm (F u)/norm u) (norm (F v)/norm v) * norm x" proof cases assume [simp]: "x \ 0" from assms have [simp]: "u \ 0" "v \ 0" by auto interpret bounded_linear F by fact define u1 where "u1 = u /\<^sub>R norm u" define v1 where "v1 = v /\<^sub>R norm v" note \x \ conefield u v\ also have \conefield u v = conefield u1 v1\ by (auto simp: u1_def v1_def conefield_scaleR conefield_commute[of u]) finally obtain c t where x: "x = c *\<^sub>R conemem u1 v1 t" "t \ {0 .. pi / 2}" "c \ 0" by (auto simp: conefield_def cone_hull_expl conesegment_def) then have xc: "x /\<^sub>R c = conemem u1 v1 t" by (auto simp: divide_simps) also have "F \ = conemem (F u1) (F v1) t" by (simp add: bounded_linear_image_conemem assms) also have "norm \ \ min (norm (F u1)) (norm (F v1)) * norm (conemem u1 v1 t)" apply (rule conemem_expansion_estimate) subgoal by fact subgoal using angle_pos by (simp add: u1_def v1_def vangle_scaleR) subgoal using angle_pos by (simp add: u1_def v1_def vangle_scaleR) subgoal using angle_le by (simp add: u1_def v1_def scaleR vangle_scaleR) subgoal using angle_le by (simp add: u1_def v1_def scaleR vangle_scaleR) subgoal using angle_le by (simp add: u1_def v1_def scaleR vangle_scaleR) done finally show "norm (F x) \ min (norm (F u)/norm u) (norm (F v)/norm v) * norm x" unfolding xc[symmetric] scaleR u1_def v1_def norm_scaleR x using \c \ 0\ by (simp add: divide_simps split: if_splits) qed simp lemma conefield_rightI: assumes ij: "i \ Basis" "j \ Basis" and ij_neq: "i \ j" assumes "y \ {y1 .. y2}" shows "(i + y *\<^sub>R j) \ conefield (i + y1 *\<^sub>R j) (i + y2 *\<^sub>R j)" unfolding conefield_alt_def apply (rule hull_inc) using assms by (auto simp: in_segment divide_simps inner_Basis algebra_simps intro!: exI[where x="(y - y1) / (y2 - y1)"] euclidean_eqI[where 'a='a] ) lemma conefield_right_vangleI: assumes ij: "i \ Basis" "j \ Basis" and ij_neq: "i \ j" assumes "y \ {y1 .. y2}" "y1 < y2" shows "(i + y *\<^sub>R j) \ conefield (i + y1 *\<^sub>R j) (i + y2 *\<^sub>R j)" unfolding conefield_alt_def apply (rule hull_inc) using assms by (auto simp: in_segment divide_simps inner_Basis algebra_simps intro!: exI[where x="(y - y1) / (y2 - y1)"] euclidean_eqI[where 'a='a] ) lemma cone_conefield[intro, simp]: "cone (conefield x y)" unfolding conefield_def by (rule cone_cone_hull) lemma conefield_mk_rightI: assumes ij: "i \ Basis" "j \ Basis" and ij_neq: "i \ j" assumes "(i + (y / x) *\<^sub>R j) \ conefield (i + (y1 / x1) *\<^sub>R j) (i + (y2 / x2) *\<^sub>R j)" assumes "x > 0" "x1 > 0" "x2 > 0" shows "(x *\<^sub>R i + y *\<^sub>R j) \ conefield (x1 *\<^sub>R i + y1 *\<^sub>R j) (x2 *\<^sub>R i + y2 *\<^sub>R j)" proof - have rescale: "(x *\<^sub>R i + y *\<^sub>R j) = x *\<^sub>R (i + (y / x) *\<^sub>R j)" if "x > 0" for x y using that by (auto simp: algebra_simps) show ?thesis unfolding rescale[OF \x > 0\] rescale[OF \x1 > 0\] rescale[OF \x2 > 0\] conefield_scaleR[OF \x1 > 0\] apply (subst conefield_commute) unfolding conefield_scaleR[OF \x2 > 0\] apply (rule mem_cone) apply simp apply (subst conefield_commute) by (auto intro!: assms less_imp_le) qed lemma conefield_prod3I: assumes "x > 0" "x1 > 0" "x2 > 0" assumes "y1 / x1 \ y / x" "y / x \ y2 / x2" shows "(x, y, 0) \ (conefield (x1, y1, 0) (x2, y2, 0)::(real*real*real) set)" proof - have "(x *\<^sub>R (1, 0, 0) + y *\<^sub>R (0, 1, 0)) \ (conefield (x1 *\<^sub>R (1, 0, 0) + y1 *\<^sub>R (0, 1, 0)) (x2 *\<^sub>R (1, 0, 0) + y2 *\<^sub>R (0, 1, 0))::(real*real*real) set)" apply (rule conefield_mk_rightI) subgoal by (auto simp: Basis_prod_def zero_prod_def) subgoal by (auto simp: Basis_prod_def zero_prod_def) subgoal by (auto simp: Basis_prod_def zero_prod_def) subgoal using assms by (intro conefield_rightI) (auto simp: Basis_prod_def zero_prod_def) by (auto intro: assms) then show ?thesis by simp qed end diff --git a/thys/Polynomial_Interpolation/Missing_Unsorted.thy b/thys/Polynomial_Interpolation/Missing_Unsorted.thy --- a/thys/Polynomial_Interpolation/Missing_Unsorted.thy +++ b/thys/Polynomial_Interpolation/Missing_Unsorted.thy @@ -1,652 +1,652 @@ (* Author: René Thiemann Akihisa Yamada Jose Divason License: BSD *) section \Missing Unsorted\ text \This theory contains several lemmas which might be of interest to the Isabelle distribution. For instance, we prove that $b^n \cdot n^k$ is bounded by a constant whenever $0 < b < 1$.\ theory Missing_Unsorted imports HOL.Complex "HOL-Computational_Algebra.Factorial_Ring" begin lemma bernoulli_inequality: assumes x: "-1 \ (x :: 'a :: linordered_field)" shows "1 + of_nat n * x \ (1 + x) ^ n" proof (induct n) case (Suc n) have "1 + of_nat (Suc n) * x = 1 + x + of_nat n * x" by (simp add: field_simps) also have "\ \ \ + of_nat n * x ^ 2" by simp also have "\ = (1 + of_nat n * x) * (1 + x)" by (simp add: field_simps power2_eq_square) also have "\ \ (1 + x) ^ n * (1 + x)" by (rule mult_right_mono[OF Suc], insert x, auto) also have "\ = (1 + x) ^ (Suc n)" by simp finally show ?case . qed simp context fixes b :: "'a :: archimedean_field" assumes b: "0 < b" "b < 1" begin private lemma pow_one: "b ^ x \ 1" using power_Suc_less_one[OF b, of "x - 1"] by (cases x, auto) private lemma pow_zero: "0 < b ^ x" using b(1) by simp lemma exp_tends_to_zero: assumes c: "c > 0" shows "\ x. b ^ x \ c" proof (rule ccontr) assume not: "\ ?thesis" define bb where "bb = inverse b" define cc where "cc = inverse c" from b have bb: "bb > 1" unfolding bb_def by (rule one_less_inverse) from c have cc: "cc > 0" unfolding cc_def by simp define bbb where "bbb = bb - 1" have id: "bb = 1 + bbb" and bbb: "bbb > 0" and bm1: "bbb \ -1" unfolding bbb_def using bb by auto have "\ n. cc / bbb < of_nat n" by (rule reals_Archimedean2) then obtain n where lt: "cc / bbb < of_nat n" by auto from not have "\ b ^ n \ c" by auto hence bnc: "b ^ n > c" by simp have "bb ^ n = inverse (b ^ n)" unfolding bb_def by (rule power_inverse) also have "\ < cc" unfolding cc_def by (rule less_imp_inverse_less[OF bnc c]) also have "\ < bbb * of_nat n" using lt bbb by (metis mult.commute pos_divide_less_eq) also have "\ \ bb ^ n" using bernoulli_inequality[OF bm1, folded id, of n] by (simp add: ac_simps) finally show False by simp qed lemma linear_exp_bound: "\ p. \ x. b ^ x * of_nat x \ p" proof - from b have "1 - b > 0" by simp from exp_tends_to_zero[OF this] obtain x0 where x0: "b ^ x0 \ 1 - b" .. { fix x assume "x \ x0" hence "\ y. x = x0 + y" by arith then obtain y where x: "x = x0 + y" by auto have "b ^ x = b ^ x0 * b ^ y" unfolding x by (simp add: power_add) also have "\ \ b ^ x0" using pow_one[of y] pow_zero[of x0] by auto also have "\ \ 1 - b" by (rule x0) finally have "b ^ x \ 1 - b" . } note x0 = this define bs where "bs = insert 1 { b ^ Suc x * of_nat (Suc x) | x . x \ x0}" have bs: "finite bs" unfolding bs_def by auto define p where "p = Max bs" have bs: "\ b. b \ bs \ b \ p" unfolding p_def using bs by simp hence p1: "p \ 1" unfolding bs_def by auto show ?thesis proof (rule exI[of _ p], intro allI) fix x show "b ^ x * of_nat x \ p" proof (induct x) case (Suc x) show ?case proof (cases "x \ x0") case True show ?thesis by (rule bs, unfold bs_def, insert True, auto) next case False let ?x = "of_nat x :: 'a" have "b ^ (Suc x) * of_nat (Suc x) = b * (b ^ x * ?x) + b ^ Suc x" by (simp add: field_simps) also have "\ \ b * p + b ^ Suc x" by (rule add_right_mono[OF mult_left_mono[OF Suc]], insert b, auto) also have "\ = p - ((1 - b) * p - b ^ (Suc x))" by (simp add: field_simps) also have "\ \ p - 0" proof - have "b ^ Suc x \ 1 - b" using x0[of "Suc x"] False by auto also have "\ \ (1 - b) * p" using b p1 by auto finally show ?thesis by (intro diff_left_mono, simp) qed finally show ?thesis by simp qed qed (insert p1, auto) qed qed lemma poly_exp_bound: "\ p. \ x. b ^ x * of_nat x ^ deg \ p" proof - show ?thesis proof (induct deg) case 0 show ?case by (rule exI[of _ 1], intro allI, insert pow_one, auto) next case (Suc deg) then obtain q where IH: "\ x. b ^ x * (of_nat x) ^ deg \ q" by auto define p where "p = max 0 q" from IH have IH: "\ x. b ^ x * (of_nat x) ^ deg \ p" unfolding p_def using le_max_iff_disj by blast have p: "p \ 0" unfolding p_def by simp show ?case proof (cases "deg = 0") case True thus ?thesis using linear_exp_bound by simp next case False note deg = this define p' where "p' = p*p * 2 ^ Suc deg * inverse b" let ?f = "\ x. b ^ x * (of_nat x) ^ Suc deg" define f where "f = ?f" { fix x let ?x = "of_nat x :: 'a" have "f (2 * x) \ (2 ^ Suc deg) * (p * p)" proof (cases "x = 0") case False hence x1: "?x \ 1" by (cases x, auto) from x1 have x: "?x ^ (deg - 1) \ 1" by simp from x1 have xx: "?x ^ Suc deg \ 1" by (rule one_le_power) define c where "c = b ^ x * b ^ x * (2 ^ Suc deg)" have c: "c > 0" unfolding c_def using b by auto have "f (2 * x) = ?f (2 * x)" unfolding f_def by simp also have "b ^ (2 * x) = (b ^ x) * (b ^ x)" by (simp add: power2_eq_square power_even_eq) also have "of_nat (2 * x) = 2 * ?x" by simp also have "(2 * ?x) ^ Suc deg = 2 ^ Suc deg * ?x ^ Suc deg" by simp finally have "f (2 * x) = c * ?x ^ Suc deg" unfolding c_def by (simp add: ac_simps) also have "\ \ c * ?x ^ Suc deg * ?x ^ (deg - 1)" proof - have "c * ?x ^ Suc deg > 0" using c xx by simp thus ?thesis unfolding mult_le_cancel_left1 using x by simp qed also have "\ = c * ?x ^ (Suc deg + (deg - 1))" by (simp add: power_add) also have "Suc deg + (deg - 1) = deg + deg" using deg by simp also have "?x ^ (deg + deg) = (?x ^ deg) * (?x ^ deg)" by (simp add: power_add) also have "c * \ = (2 ^ Suc deg) * ((b ^ x * ?x ^ deg) * (b ^ x * ?x ^ deg))" unfolding c_def by (simp add: ac_simps) also have "\ \ (2 ^ Suc deg) * (p * p)" by (rule mult_left_mono[OF mult_mono[OF IH IH p]], insert pow_zero[of x], auto) finally show "f (2 * x) \ (2 ^ Suc deg) * (p * p)" . qed (auto simp: f_def) hence "?f (2 * x) \ (2 ^ Suc deg) * (p * p)" unfolding f_def . } note even = this show ?thesis proof (rule exI[of _ p'], intro allI) fix y show "?f y \ p'" proof (cases "even y") case True define x where "x = y div 2" have "y = 2 * x" unfolding x_def using True by simp from even[of x, folded this] have "?f y \ 2 ^ Suc deg * (p * p)" . also have "\ \ \ * inverse b" unfolding mult_le_cancel_left1 using b p - by (simp add: sign_simps one_le_inverse) + by (simp add: algebra_split_simps one_le_inverse) also have "\ = p'" unfolding p'_def by (simp add: ac_simps) finally show "?f y \ p'" . next case False define x where "x = y div 2" have "y = 2 * x + 1" unfolding x_def using False by simp hence "?f y = ?f (2 * x + 1)" by simp also have "\ \ b ^ (2 * x + 1) * of_nat (2 * x + 2) ^ Suc deg" by (rule mult_left_mono[OF power_mono], insert b, auto) also have "b ^ (2 * x + 1) = b ^ (2 * x + 2) * inverse b" using b by auto also have "b ^ (2 * x + 2) * inverse b * of_nat (2 * x + 2) ^ Suc deg = inverse b * ?f (2 * (x + 1))" by (simp add: ac_simps) also have "\ \ inverse b * ((2 ^ Suc deg) * (p * p))" by (rule mult_left_mono[OF even], insert b, auto) also have "\ = p'" unfolding p'_def by (simp add: ac_simps) finally show "?f y \ p'" . qed qed qed qed qed end lemma prod_list_replicate[simp]: "prod_list (replicate n a) = a ^ n" by (induct n, auto) lemma prod_list_power: fixes xs :: "'a :: comm_monoid_mult list" shows "prod_list xs ^ n = (\x\xs. x ^ n)" by (induct xs, auto simp: power_mult_distrib) lemma set_upt_Suc: "{0 ..< Suc i} = insert i {0 ..< i}" by (fact atLeast0_lessThan_Suc) lemma prod_pow[simp]: "(\i = 0..a\ * y dvd x \ a * y dvd x" for x y a :: int using abs_dvd_iff [of "a * y"] abs_dvd_iff [of "\a\ * y"] by (simp add: abs_mult) lemma gcd_abs_mult_right_int [simp]: "gcd x (\a\ * y) = gcd x (a * y)" for x y a :: int using gcd_abs2_int [of _ "a * y"] gcd_abs2_int [of _ "\a\ * y"] by (simp add: abs_mult) lemma lcm_abs_mult_right_int [simp]: "lcm x (\a\ * y) = lcm x (a * y)" for x y a :: int using lcm_abs2_int [of _ "a * y"] lcm_abs2_int [of _ "\a\ * y"] by (simp add: abs_mult) lemma gcd_abs_mult_left_int [simp]: "gcd x (a * \y\) = gcd x (a * y)" for x y a :: int using gcd_abs2_int [of _ "a * \y\"] gcd_abs2_int [of _ "a * y"] by (simp add: abs_mult) lemma lcm_abs_mult_left_int [simp]: "lcm x (a * \y\) = lcm x (a * y)" for x y a :: int using lcm_abs2_int [of _ "a * \y\"] lcm_abs2_int [of _ "a * y"] by (simp add: abs_mult) abbreviation (input) list_gcd :: "'a :: semiring_gcd list \ 'a" where "list_gcd \ gcd_list" abbreviation (input) list_lcm :: "'a :: semiring_gcd list \ 'a" where "list_lcm \ lcm_list" lemma list_gcd_simps: "list_gcd [] = 0" "list_gcd (x # xs) = gcd x (list_gcd xs)" by simp_all lemma list_gcd: "x \ set xs \ list_gcd xs dvd x" by (fact Gcd_fin_dvd) lemma list_gcd_greatest: "(\ x. x \ set xs \ y dvd x) \ y dvd (list_gcd xs)" by (fact gcd_list_greatest) lemma list_gcd_mult_int [simp]: fixes xs :: "int list" shows "list_gcd (map (times a) xs) = \a\ * list_gcd xs" by (simp add: Gcd_mult) lemma list_lcm_simps: "list_lcm [] = 1" "list_lcm (x # xs) = lcm x (list_lcm xs)" by simp_all lemma list_lcm: "x \ set xs \ x dvd list_lcm xs" by (fact dvd_Lcm_fin) lemma list_lcm_least: "(\ x. x \ set xs \ x dvd y) \ list_lcm xs dvd y" by (fact lcm_list_least) lemma lcm_mult_distrib_nat: "(k :: nat) * lcm m n = lcm (k * m) (k * n)" by (simp add: lcm_mult_left) lemma lcm_mult_distrib_int: "abs (k::int) * lcm m n = lcm (k * m) (k * n)" by (simp add: lcm_mult_left) lemma list_lcm_mult_int [simp]: fixes xs :: "int list" shows "list_lcm (map (times a) xs) = (if xs = [] then 1 else \a\ * list_lcm xs)" by (simp add: Lcm_mult) lemma list_lcm_pos: "list_lcm xs \ (0 :: int)" "0 \ set xs \ list_lcm xs \ 0" "0 \ set xs \ list_lcm xs > 0" proof - have "0 \ \Lcm (set xs)\" by (simp only: abs_ge_zero) then have "0 \ Lcm (set xs)" by simp then show "list_lcm xs \ 0" by simp assume "0 \ set xs" then show "list_lcm xs \ 0" by (simp add: Lcm_0_iff) with \list_lcm xs \ 0\ show "list_lcm xs > 0" by (simp add: le_less) qed lemma quotient_of_nonzero: "snd (quotient_of r) > 0" "snd (quotient_of r) \ 0" using quotient_of_denom_pos' [of r] by simp_all lemma quotient_of_int_div: assumes q: "quotient_of (of_int x / of_int y) = (a, b)" and y: "y \ 0" shows "\ z. z \ 0 \ x = a * z \ y = b * z" proof - let ?r = "rat_of_int" define z where "z = gcd x y" define x' where "x' = x div z" define y' where "y' = y div z" have id: "x = z * x'" "y = z * y'" unfolding x'_def y'_def z_def by auto from y have y': "y' \ 0" unfolding id by auto have z: "z \ 0" unfolding z_def using y by auto have cop: "coprime x' y'" unfolding x'_def y'_def z_def using div_gcd_coprime y by blast have "?r x / ?r y = ?r x' / ?r y'" unfolding id using z y y' by (auto simp: field_simps) from assms[unfolded this] have quot: "quotient_of (?r x' / ?r y') = (a, b)" by auto from quotient_of_coprime[OF quot] have cop': "coprime a b" . hence cop: "coprime b a" by (simp add: ac_simps) from quotient_of_denom_pos[OF quot] have b: "b > 0" "b \ 0" by auto from quotient_of_div[OF quot] quotient_of_denom_pos[OF quot] y' have "?r x' * ?r b = ?r a * ?r y'" by (auto simp: field_simps) hence id': "x' * b = a * y'" unfolding of_int_mult[symmetric] by linarith from id'[symmetric] have "b dvd y' * a" unfolding mult.commute[of y'] by auto with cop y' have "b dvd y'" by (simp add: coprime_dvd_mult_left_iff) then obtain z' where ybz: "y' = b * z'" unfolding dvd_def by auto from id[unfolded y' this] have y: "y = b * (z * z')" by auto with \y \ 0\ have zz: "z * z' \ 0" by auto from quotient_of_div[OF q] \y \ 0\ \b \ 0\ have "?r x * ?r b = ?r y * ?r a" by (auto simp: field_simps) hence id': "x * b = y * a" unfolding of_int_mult[symmetric] by linarith from this[unfolded y] b have x: "x = a * (z * z')" by auto show ?thesis unfolding x y using zz by blast qed fun max_list_non_empty :: "('a :: linorder) list \ 'a" where "max_list_non_empty [x] = x" | "max_list_non_empty (x # xs) = max x (max_list_non_empty xs)" lemma max_list_non_empty: "x \ set xs \ x \ max_list_non_empty xs" proof (induct xs) case (Cons y ys) note oCons = this show ?case proof (cases ys) case (Cons z zs) hence id: "max_list_non_empty (y # ys) = max y (max_list_non_empty ys)" by simp from oCons show ?thesis unfolding id by (auto simp: max.coboundedI2) qed (insert oCons, auto) qed simp lemma cnj_reals[simp]: "(cnj c \ \) = (c \ \)" using Reals_cnj_iff by fastforce lemma sgn_real_mono: "x \ y \ sgn x \ sgn (y :: real)" unfolding sgn_real_def by (auto split: if_splits) lemma sgn_minus_rat: "sgn (- (x :: rat)) = - sgn x" by (fact Rings.sgn_minus) lemma real_of_rat_sgn: "sgn (of_rat x) = real_of_rat (sgn x)" unfolding sgn_real_def sgn_rat_def by auto lemma inverse_le_iff_sgn: assumes sgn: "sgn x = sgn y" shows "(inverse (x :: real) \ inverse y) = (y \ x)" proof (cases "x = 0") case True with sgn have "sgn y = 0" by simp hence "y = 0" unfolding sgn_real_def by (cases "y = 0"; cases "y < 0"; auto) thus ?thesis using True by simp next case False note x = this show ?thesis proof (cases "x < 0") case True with x sgn have "sgn y = -1" by simp hence "y < 0" unfolding sgn_real_def by (cases "y = 0"; cases "y < 0", auto) show ?thesis by (rule inverse_le_iff_le_neg[OF True \y < 0\]) next case False with x have x: "x > 0" by auto with sgn have "sgn y = 1" by auto hence "y > 0" unfolding sgn_real_def by (cases "y = 0"; cases "y < 0", auto) show ?thesis by (rule inverse_le_iff_le[OF x \y > 0\]) qed qed lemma inverse_le_sgn: assumes sgn: "sgn x = sgn y" and xy: "x \ (y :: real)" shows "inverse y \ inverse x" using xy inverse_le_iff_sgn[OF sgn] by auto lemma set_list_update: "set (xs [i := k]) = (if i < length xs then insert k (set (take i xs) \ set (drop (Suc i) xs)) else set xs)" proof (induct xs arbitrary: i) case (Cons x xs i) thus ?case by (cases i, auto) qed simp lemma prod_list_dvd: assumes "(x :: 'a :: comm_monoid_mult) \ set xs" shows "x dvd prod_list xs" proof - from assms[unfolded in_set_conv_decomp] obtain ys zs where xs: "xs = ys @ x # zs" by auto show ?thesis unfolding xs dvd_def by (intro exI[of _ "prod_list (ys @ zs)"], simp add: ac_simps) qed lemma dvd_prod: fixes A::"'b set" assumes "\b\A. a dvd f b" "finite A" shows "a dvd prod f A" using assms(2,1) proof (induct A) case (insert x A) thus ?case using comm_monoid_mult_class.dvd_mult dvd_mult2 insert_iff prod.insert by auto qed auto context fixes xs :: "'a :: comm_monoid_mult list" begin lemma prod_list_filter: "prod_list (filter f xs) * prod_list (filter (\ x. \ f x) xs) = prod_list xs" by (induct xs, auto simp: ac_simps) lemma prod_list_partition: assumes "partition f xs = (ys, zs)" shows "prod_list xs = prod_list ys * prod_list zs" using assms by (subst prod_list_filter[symmetric, of f], auto simp: o_def) end lemma dvd_imp_mult_div_cancel_left[simp]: assumes "(a :: 'a :: semidom_divide) dvd b" shows "a * (b div a) = b" proof(cases "b = 0") case True then show ?thesis by auto next case False with dvdE[OF assms] obtain c where *: "b = a * c" by auto also with False have "a \ 0" by auto then have "a * c div a = c" by auto also note *[symmetric] finally show ?thesis. qed lemma (in semidom) prod_list_zero_iff[simp]: "prod_list xs = 0 \ 0 \ set xs" by (induction xs, auto) context comm_monoid_mult begin lemma unit_prod [intro]: shows "a dvd 1 \ b dvd 1 \ (a * b) dvd 1" by (subst mult_1_left [of 1, symmetric]) (rule mult_dvd_mono) lemma is_unit_mult_iff[simp]: shows "(a * b) dvd 1 \ a dvd 1 \ b dvd 1" by (auto dest: dvd_mult_left dvd_mult_right) end context comm_semiring_1 begin lemma irreducibleE[elim]: assumes "irreducible p" and "p \ 0 \ \ p dvd 1 \ (\a b. p = a * b \ a dvd 1 \ b dvd 1) \ thesis" shows thesis using assms by (auto simp: irreducible_def) lemma not_irreducibleE: assumes "\ irreducible x" and "x = 0 \ thesis" and "x dvd 1 \ thesis" and "\a b. x = a * b \ \ a dvd 1 \ \ b dvd 1 \ thesis" shows thesis using assms unfolding irreducible_def by auto lemma prime_elem_dvd_prod_list: assumes p: "prime_elem p" and pA: "p dvd prod_list A" shows "\a \ set A. p dvd a" proof(insert pA, induct A) case Nil with p show ?case by (simp add: prime_elem_not_unit) next case (Cons a A) then show ?case by (auto simp: prime_elem_dvd_mult_iff[OF p]) qed lemma prime_elem_dvd_prod_mset: assumes p: "prime_elem p" and pA: "p dvd prod_mset A" shows "\a \# A. p dvd a" proof(insert pA, induct A) case empty with p show ?case by (simp add: prime_elem_not_unit) next case (add a A) then show ?case by (auto simp: prime_elem_dvd_mult_iff[OF p]) qed lemma mult_unit_dvd_iff[simp]: assumes "b dvd 1" shows "a * b dvd c \ a dvd c" proof assume "a * b dvd c" with assms show "a dvd c" using dvd_mult_left[of a b c] by simp next assume "a dvd c" with assms mult_dvd_mono show "a * b dvd c" by fastforce qed lemma mult_unit_dvd_iff'[simp]: "a dvd 1 \ (a * b) dvd c \ b dvd c" using mult_unit_dvd_iff [of a b c] by (simp add: ac_simps) lemma irreducibleD': assumes "irreducible a" "b dvd a" shows "a dvd b \ b dvd 1" proof - from assms obtain c where c: "a = b * c" by (elim dvdE) from irreducibleD[OF assms(1) this] have "b dvd 1 \ c dvd 1" . thus ?thesis by (auto simp: c) qed end context idom begin text \ Following lemmas are adapted and generalized so that they don't use "algebraic" classes. \ lemma dvd_times_left_cancel_iff [simp]: assumes "a \ 0" shows "a * b dvd a * c \ b dvd c" (is "?lhs \ ?rhs") proof assume ?lhs then obtain d where "a * c = a * b * d" .. with assms have "c = b * d" by (auto simp add: ac_simps) then show ?rhs .. next assume ?rhs then obtain d where "c = b * d" .. then have "a * c = a * b * d" by (simp add: ac_simps) then show ?lhs .. qed lemma dvd_times_right_cancel_iff [simp]: assumes "a \ 0" shows "b * a dvd c * a \ b dvd c" using dvd_times_left_cancel_iff [of a b c] assms by (simp add: ac_simps) lemma irreducibleI': assumes "a \ 0" "\ a dvd 1" "\b. b dvd a \ a dvd b \ b dvd 1" shows "irreducible a" proof (rule irreducibleI) fix b c assume a_eq: "a = b * c" hence "a dvd b \ b dvd 1" by (intro assms) simp_all thus "b dvd 1 \ c dvd 1" proof assume "a dvd b" hence "b * c dvd b * 1" by (simp add: a_eq) moreover from \a \ 0\ a_eq have "b \ 0" by auto ultimately show ?thesis using dvd_times_left_cancel_iff by fastforce qed blast qed (simp_all add: assms(1,2)) lemma irreducible_altdef: shows "irreducible x \ x \ 0 \ \ x dvd 1 \ (\b. b dvd x \ x dvd b \ b dvd 1)" using irreducibleI'[of x] irreducibleD'[of x] irreducible_not_unit[of x] by auto lemma dvd_mult_unit_iff: assumes b: "b dvd 1" shows "a dvd c * b \ a dvd c" proof- from b obtain b' where 1: "b * b' = 1" by (elim dvdE, auto) then have b0: "b \ 0" by auto from 1 have "a = (a * b') * b" by (simp add: ac_simps) also have "\ dvd c * b \ a * b' dvd c" using b0 by auto finally show ?thesis by (auto intro: dvd_mult_left) qed lemma dvd_mult_unit_iff': "b dvd 1 \ a dvd b * c \ a dvd c" using dvd_mult_unit_iff [of b a c] by (simp add: ac_simps) lemma irreducible_mult_unit_left: shows "a dvd 1 \ irreducible (a * p) \ irreducible p" by (auto simp: irreducible_altdef mult.commute[of a] dvd_mult_unit_iff) lemma irreducible_mult_unit_right: shows "a dvd 1 \ irreducible (p * a) \ irreducible p" by (auto simp: irreducible_altdef mult.commute[of a] dvd_mult_unit_iff) lemma prime_elem_imp_irreducible: assumes "prime_elem p" shows "irreducible p" proof (rule irreducibleI) fix a b assume p_eq: "p = a * b" with assms have nz: "a \ 0" "b \ 0" by auto from p_eq have "p dvd a * b" by simp with \prime_elem p\ have "p dvd a \ p dvd b" by (rule prime_elem_dvd_multD) with \p = a * b\ have "a * b dvd 1 * b \ a * b dvd a * 1" by auto thus "a dvd 1 \ b dvd 1" by (simp only: dvd_times_left_cancel_iff[OF nz(1)] dvd_times_right_cancel_iff[OF nz(2)]) qed (insert assms, simp_all add: prime_elem_def) lemma unit_imp_dvd [dest]: "b dvd 1 \ b dvd a" by (rule dvd_trans [of _ 1]) simp_all lemma unit_mult_left_cancel: "a dvd 1 \ a * b = a * c \ b = c" using mult_cancel_left [of a b c] by auto lemma unit_mult_right_cancel: "a dvd 1 \ b * a = c * a \ b = c" using unit_mult_left_cancel [of a b c] by (auto simp add: ac_simps) text \New parts from here\ lemma irreducible_multD: assumes l: "irreducible (a*b)" shows "a dvd 1 \ irreducible b \ b dvd 1 \ irreducible a" proof- from l have "a dvd 1 \ b dvd 1" using irreducibleD by auto then show ?thesis proof(elim disjE) assume a: "a dvd 1" with l have "irreducible b" unfolding irreducible_def by (metis is_unit_mult_iff mult.left_commute mult_not_zero) with a show ?thesis by auto next assume a: "b dvd 1" with l have "irreducible a" unfolding irreducible_def by (meson is_unit_mult_iff mult_not_zero semiring_normalization_rules(16)) with a show ?thesis by auto qed qed end lemma (in field) irreducible_field[simp]: "irreducible x \ False" by (auto simp: dvd_field_iff irreducible_def) lemma (in idom) irreducible_mult: shows "irreducible (a*b) \ a dvd 1 \ irreducible b \ b dvd 1 \ irreducible a" by (auto dest: irreducible_multD simp: irreducible_mult_unit_left irreducible_mult_unit_right) 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 diff --git a/thys/Priority_Search_Trees/PST_General.thy b/thys/Priority_Search_Trees/PST_General.thy --- a/thys/Priority_Search_Trees/PST_General.thy +++ b/thys/Priority_Search_Trees/PST_General.thy @@ -1,116 +1,116 @@ section \General Priority Search Trees\ theory PST_General imports "HOL-Data_Structures.Tree2" Prio_Map_Specs begin text \\noindent We show how to implement priority maps by augmented binary search trees. That is, the basic data structure is some arbitrary binary search tree, e.g.\ a red-black tree, implementing the map from @{typ 'a} to @{typ 'b} by storing pairs \(k,p)\ in each node. At this point we need to assume that the keys are also linearly ordered. To implement \getmin\ efficiently we annotate/augment each node with another pair \(k',p')\, the intended result of \getmin\ when applied to that subtree. The specification of \getmin\ tells us that \(k',p')\ must be in that subtree and that \p'\ is the minimal priority in that subtree. Thus the annotation can be computed by passing the \(k',p')\ with the minimal \p'\ up the tree. We will now make this more precise for balanced binary trees in general. We assume that our trees are either leaves of the form @{term Leaf} or nodes of the form @{term "Node l (kp, b) r"} where \l\ and \r\ are subtrees, \kp\ is the contents of the node (a key-priority pair) and \b\ is some additional balance information (e.g.\ colour, height, size, \dots). Augmented nodes are of the form \<^term>\Node l (kp, (b,kp')) r\. \ type_synonym ('k,'p,'c) pstree = "(('k\'p) \ ('c \ ('k \ 'p))) tree" text \ The following invariant states that a node annotation is actually a minimal key-priority pair for the node's subtree. \ fun invpst :: "('k,'p::linorder,'c) pstree \ bool" where "invpst Leaf = True" | "invpst (Node l (x, _,mkp) r) \ invpst l \ invpst r \ is_min2 mkp (set (inorder l @ x # inorder r))" text \The implementation of \getmin\ is trivial:\ fun pst_getmin where "pst_getmin (Node _ (_, _,a) _) = a" lemma pst_getmin_ismin: - "invpst t \ t\Leaf \ is_min2 (pst_getmin t) (set (inorder t))" + "invpst t \ t\Leaf \ is_min2 (pst_getmin t) (set_tree t)" by (cases t rule: pst_getmin.cases) auto text \ It remains to upgrade the existing map operations to work with augmented nodes. Therefore we now show how to transform any function definition on un-augmented trees into one on trees augmented with \(k',p')\ pairs. A defining equation \f pats = e\ for the original type of nodes is transformed into an equation \f pats' = e'\ on the augmented type of nodes as follows: \<^item> Every pattern @{term "Node l (kp, b) r"} in \pats\ and \e\ is replaced by @{term "Node l (kp, (b,DUMMY)) r"} to obtain \pats'\ and \e\<^sub>2\. \<^item> To obtain \e'\, every expression @{term "Node l (kp, b) r"} in \e\<^sub>2\ is replaced by \mkNode l kp b r\ where: \ definition "min2 \ \(k,p) (k',p'). if p\p' then (k,p) else (k',p')" definition "min_kp a l r \ case (l,r) of (Leaf,Leaf) \ a | (Leaf,Node _ (_, (_,kpr)) _) \ min2 a kpr | (Node _ (_, (_,kpl)) _,Leaf) \ min2 a kpl | (Node _ (_, (_,kpl)) _,Node _ (_, (_,kpr)) _) \ min2 a (min2 kpl kpr)" definition "mkNode c l a r \ Node l (a, (c,min_kp a l r)) r" text \ Note that this transformation does not affect the asymptotic complexity of \f\. Therefore the priority search tree operations have the same complexity as the underlying search tree operations, i.e.\ typically logarithmic (\update\, \delete\, \lookup\) and constant time (\empty\, \is_empty\). \ text \It is straightforward to show that @{const mkNode} preserves the invariant:\ lemma is_min2_Empty[simp]: "\is_min2 x {}" by (auto simp: is_min2_def) lemma is_min2_singleton[simp]: "is_min2 a {b} \ b=a" by (auto simp: is_min2_def) lemma is_min2_insert: "is_min2 x (insert y ys) \ (y=x \ (\z\ys. snd x \ snd z)) \ (snd x \ snd y \ is_min2 x ys)" by (auto simp: is_min2_def) lemma is_min2_union: "is_min2 x (ys \ zs) \ (is_min2 x ys \ (\z\zs. snd x \ snd z)) \ ((\y\ys. snd x \ snd y) \ is_min2 x zs)" by (auto simp: is_min2_def) lemma is_min2_min2_insI: "is_min2 y ys \ is_min2 (min2 x y) (insert x ys)" by (auto simp: is_min2_def min2_def split: prod.split) lemma is_min2_mergeI: "is_min2 x xs \ is_min2 y ys \ is_min2 (min2 x y) (xs \ ys)" by (auto simp: is_min2_def min2_def split: prod.split) theorem invpst_mkNode[simp]: "invpst (mkNode c l a r) \ invpst l \ invpst r" apply (cases l rule: invpst.cases; cases r rule: invpst.cases; simp add: mkNode_def min_kp_def) subgoal using is_min2_min2_insI by blast subgoal by (auto intro!: is_min2_min2_insI simp: insert_commute) subgoal by (smt Un_insert_left Un_insert_right is_min2_mergeI is_min2_min2_insI sup_assoc) done end diff --git a/thys/Priority_Search_Trees/PST_RBT.thy b/thys/Priority_Search_Trees/PST_RBT.thy --- a/thys/Priority_Search_Trees/PST_RBT.thy +++ b/thys/Priority_Search_Trees/PST_RBT.thy @@ -1,411 +1,411 @@ section \Priority Search Trees on top of RBTs\ theory PST_RBT imports "HOL-Data_Structures.Cmp" "HOL-Data_Structures.Isin2" "HOL-Data_Structures.Lookup2" PST_General begin text \ We obtain a priority search map based on red-black trees via the general priority search tree augmentation. This theory has been derived from the standard Isabelle implementation of red black trees in @{session "HOL-Data_Structures"}. \ subsection \Definitions\ subsubsection \The Code\ datatype tcolor = Red | Black type_synonym ('k,'p) rbth = "(('k\'p) \ (tcolor \ ('k \ 'p))) tree" abbreviation R where "R mkp l a r \ Node l (a, Red,mkp) r" abbreviation B where "B mkp l a r \ Node l (a, Black,mkp) r" abbreviation "mkR \ mkNode Red" abbreviation "mkB \ mkNode Black" fun baliL :: "('k,'p::linorder) rbth \ 'k\'p \ ('k,'p) rbth \ ('k,'p) rbth" where "baliL (R _ (R _ t1 a1 t2) a2 t3) a3 t4 = mkR (mkB t1 a1 t2) a2 (mkB t3 a3 t4)" | "baliL (R _ t1 a1 (R _ t2 a2 t3)) a3 t4 = mkR (mkB t1 a1 t2) a2 (mkB t3 a3 t4)" | "baliL t1 a t2 = mkB t1 a t2" fun baliR :: "('k,'p::linorder) rbth \ 'k\'p \ ('k,'p) rbth \ ('k,'p) rbth" where "baliR t1 a1 (R _ (R _ t2 a2 t3) a3 t4) = mkR (mkB t1 a1 t2) a2 (mkB t3 a3 t4)" | "baliR t1 a1 (R _ t2 a2 (R _ t3 a3 t4)) = mkR (mkB t1 a1 t2) a2 (mkB t3 a3 t4)" | "baliR t1 a t2 = mkB t1 a t2" fun paint :: "tcolor \ ('k,'p::linorder) rbth \ ('k,'p::linorder) rbth" where "paint c Leaf = Leaf" | "paint c (Node l (a, (_,mkp)) r) = Node l (a, (c,mkp)) r" fun baldL :: "('k,'p::linorder) rbth \ 'k \ 'p \ ('k,'p::linorder) rbth \ ('k,'p::linorder) rbth" where "baldL (R _ t1 x t2) y t3 = mkR (mkB t1 x t2) y t3" | "baldL bl x (B _ t1 y t2) = baliR bl x (mkR t1 y t2)" | "baldL bl x (R _ (B _ t1 y t2) z t3) = mkR (mkB bl x t1) y (baliR t2 z (paint Red t3))" | "baldL t1 x t2 = mkR t1 x t2" fun baldR :: "('k,'p::linorder) rbth \ 'k \ 'p \ ('k,'p::linorder) rbth \ ('k,'p::linorder) rbth" where "baldR t1 x (R _ t2 y t3) = mkR t1 x (mkB t2 y t3)" | "baldR (B _ t1 x t2) y t3 = baliL (mkR t1 x t2) y t3" | "baldR (R _ t1 x (B _ t2 y t3)) z t4 = mkR (baliL (paint Red t1) x t2) y (mkB t3 z t4)" | "baldR t1 x t2 = mkR t1 x t2" fun combine :: "('k,'p::linorder) rbth \ ('k,'p::linorder) rbth \ ('k,'p::linorder) rbth" where "combine Leaf t = t" | "combine t Leaf = t" | "combine (R _ t1 a t2) (R _ t3 c t4) = (case combine t2 t3 of R _ u2 b u3 \ (mkR (mkR t1 a u2) b (mkR u3 c t4)) | t23 \ mkR t1 a (mkR t23 c t4))" | "combine (B _ t1 a t2) (B _ t3 c t4) = (case combine t2 t3 of R _ t2' b t3' \ mkR (mkB t1 a t2') b (mkB t3' c t4) | t23 \ baldL t1 a (mkB t23 c t4))" | "combine t1 (R _ t2 a t3) = mkR (combine t1 t2) a t3" | "combine (R _ t1 a t2) t3 = mkR t1 a (combine t2 t3)" fun color :: "('k,'p) rbth \ tcolor" where "color Leaf = Black" | "color (Node _ (_, (c,_)) _) = c" fun upd :: "'a::linorder \ 'b::linorder \ ('a,'b) rbth \ ('a,'b) rbth" where "upd x y Leaf = mkR Leaf (x,y) Leaf" | "upd x y (B _ l (a,b) r) = (case cmp x a of LT \ baliL (upd x y l) (a,b) r | GT \ baliR l (a,b) (upd x y r) | EQ \ mkB l (x,y) r)" | "upd x y (R _ l (a,b) r) = (case cmp x a of LT \ mkR (upd x y l) (a,b) r | GT \ mkR l (a,b) (upd x y r) | EQ \ mkR l (x,y) r)" definition update :: "'a::linorder \ 'b::linorder \ ('a,'b) rbth \ ('a,'b) rbth" where "update x y t = paint Black (upd x y t)" fun del :: "'a::linorder \ ('a,'b::linorder)rbth \ ('a,'b)rbth" where "del x Leaf = Leaf" | "del x (Node l ((a,b), (c,_)) r) = (case cmp x a of LT \ if l \ Leaf \ color l = Black then baldL (del x l) (a,b) r else mkR (del x l) (a,b) r | GT \ if r \ Leaf\ color r = Black then baldR l (a,b) (del x r) else mkR l (a,b) (del x r) | EQ \ combine l r)" definition delete :: "'a::linorder \ ('a,'b::linorder) rbth \ ('a,'b) rbth" where "delete x t = paint Black (del x t)" subsubsection \Invariants\ fun bheight :: "('k,'p) rbth \ nat" where "bheight Leaf = 0" | "bheight (Node l (x, (c,_)) r) = (if c = Black then bheight l + 1 else bheight l)" fun invc :: "('k,'p) rbth \ bool" where "invc Leaf = True" | "invc (Node l (a, (c,_)) r) = (invc l \ invc r \ (c = Red \ color l = Black \ color r = Black))" fun invc2 :: "('k,'p) rbth \ bool" \ \Weaker version\ where "invc2 Leaf = True" | "invc2 (Node l (a, _) r) = (invc l \ invc r)" fun invh :: "('k,'p) rbth \ bool" where "invh Leaf = True" | "invh (Node l (x, _) r) = (invh l \ invh r \ bheight l = bheight r)" definition rbt :: "('k,'p::linorder) rbth \ bool" where "rbt t = (invc t \ invh t \ invpst t \ color t = Black)" subsection \Functional Correctness\ lemma inorder_paint[simp]: "inorder(paint c t) = inorder t" by(cases t) (auto) lemma inorder_mkNode[simp]: "inorder (mkNode c l a r) = inorder l @ a # inorder r" by (auto simp: mkNode_def) lemma inorder_baliL[simp]: "inorder(baliL l a r) = inorder l @ a # inorder r" by(cases "(l,a,r)" rule: baliL.cases) (auto) lemma inorder_baliR[simp]: "inorder(baliR l a r) = inorder l @ a # inorder r" by(cases "(l,a,r)" rule: baliR.cases) (auto) lemma inorder_baldL[simp]: "inorder(baldL l a r) = inorder l @ a # inorder r" by (cases "(l,a,r)" rule: baldL.cases) auto lemma inorder_baldR[simp]: "inorder(baldR l a r) = inorder l @ a # inorder r" by(cases "(l,a,r)" rule: baldR.cases) auto lemma inorder_combine[simp]: "inorder(combine l r) = inorder l @ inorder r" by (induction l r rule: combine.induct) (auto split: tree.split tcolor.split) lemma inorder_upd: "sorted1(inorder t) \ inorder(upd x y t) = upd_list x y (inorder t)" by(induction x y t rule: upd.induct) (auto simp: upd_list_simps) lemma inorder_update: "sorted1(inorder t) \ inorder(update x y t) = upd_list x y (inorder t)" by(simp add: update_def inorder_upd) lemma inorder_del: "sorted1(inorder t) \ inorder(del x t) = del_list x (inorder t)" by(induction x t rule: del.induct) (auto simp: del_list_simps) lemma inorder_delete: "sorted1(inorder t) \ inorder(delete x t) = del_list x (inorder t)" by(simp add: delete_def inorder_del) subsection \Invariant Preservation\ lemma color_paint_Black: "color (paint Black t) = Black" by (cases t) auto theorem rbt_Leaf: "rbt Leaf" by (simp add: rbt_def) lemma invc2I: "invc t \ invc2 t" by (cases t rule: invc.cases) simp+ lemma paint_invc2: "invc2 t \ invc2 (paint c t)" by (cases t) auto lemma invc_paint_Black: "invc2 t \ invc (paint Black t)" by (cases t) auto lemma invh_paint: "invh t \ invh (paint c t)" by (cases t) auto lemma invc_mkRB[simp]: "invc (mkR l a r) \ invc l \ invc r \ color l = Black \ color r = Black" "invc (mkB l a r) \ invc l \ invc r" by (simp_all add: mkNode_def) lemma color_mkNode[simp]: "color (mkNode c l a r) = c" by (simp_all add: mkNode_def) subsubsection \Update\ lemma invc_baliL: "\invc2 l; invc r\ \ invc (baliL l a r)" by (induct l a r rule: baliL.induct) auto lemma invc_baliR: "\invc l; invc2 r\ \ invc (baliR l a r)" by (induct l a r rule: baliR.induct) auto lemma bheight_mkRB[simp]: "bheight (mkR l a r) = bheight l" "bheight (mkB l a r) = Suc (bheight l)" by (simp_all add: mkNode_def) lemma bheight_baliL: "bheight l = bheight r \ bheight (baliL l a r) = Suc (bheight l)" by (induct l a r rule: baliL.induct) auto lemma bheight_baliR: "bheight l = bheight r \ bheight (baliR l a r) = Suc (bheight l)" by (induct l a r rule: baliR.induct) auto lemma invh_mkNode[simp]: "invh (mkNode c l a r) \ invh l \ invh r \ bheight l = bheight r" by (simp add: mkNode_def) lemma invh_baliL: "\ invh l; invh r; bheight l = bheight r \ \ invh (baliL l a r)" by (induct l a r rule: baliL.induct) auto lemma invh_baliR: "\ invh l; invh r; bheight l = bheight r \ \ invh (baliR l a r)" by (induct l a r rule: baliR.induct) auto lemma invc_upd: assumes "invc t" shows "color t = Black \ invc (upd x y t)" "invc2 (upd x y t)" using assms by (induct x y t rule: upd.induct) (auto simp: invc_baliL invc_baliR invc2I mkNode_def) lemma invh_upd: assumes "invh t" shows "invh (upd x y t)" "bheight (upd x y t) = bheight t" using assms by(induct x y t rule: upd.induct) (auto simp: invh_baliL invh_baliR bheight_baliL bheight_baliR) lemma invpst_paint[simp]: "invpst (paint c t) = invpst t" by (cases "(c,t)" rule: paint.cases) auto lemma invpst_baliR: "invpst l \ invpst r \ invpst (baliR l a r)" by (cases "(l,a,r)" rule: baliR.cases) auto lemma invpst_baliL: "invpst l \ invpst r \ invpst (baliL l a r)" by (cases "(l,a,r)" rule: baliL.cases) auto lemma invpst_upd: "invpst t \ invpst (upd x y t)" by (induct x y t rule: upd.induct) (auto simp: invpst_baliR invpst_baliL) theorem rbt_update: "rbt t \ rbt (update x y t)" by (simp add: invc_upd(2) invh_upd(1) color_paint_Black invc_paint_Black invh_paint rbt_def update_def invpst_upd) subsubsection \Delete\ lemma bheight_paint_Red: "color t = Black \ bheight (paint Red t) = bheight t - 1" by (cases t) auto lemma invh_baldL_invc: "\ invh l; invh r; bheight l + 1 = bheight r; invc r \ \ invh (baldL l a r) \ bheight (baldL l a r) = bheight l + 1" by (induct l a r rule: baldL.induct) (auto simp: invh_baliR invh_paint bheight_baliR bheight_paint_Red) lemma invh_baldL_Black: "\ invh l; invh r; bheight l + 1 = bheight r; color r = Black \ \ invh (baldL l a r) \ bheight (baldL l a r) = bheight r" by (induct l a r rule: baldL.induct) (auto simp add: invh_baliR bheight_baliR) lemma invc_baldL: "\invc2 l; invc r; color r = Black\ \ invc (baldL l a r)" by (induct l a r rule: baldL.induct) (auto simp: invc_baliR invc2I mkNode_def) lemma invc2_baldL: "\ invc2 l; invc r \ \ invc2 (baldL l a r)" by (induct l a r rule: baldL.induct) (auto simp: invc_baliR paint_invc2 invc2I mkNode_def) lemma invh_baldR_invc: "\ invh l; invh r; bheight l = bheight r + 1; invc l \ \ invh (baldR l a r) \ bheight (baldR l a r) = bheight l" by(induct l a r rule: baldR.induct) (auto simp: invh_baliL bheight_baliL invh_paint bheight_paint_Red) lemma invc_baldR: "\invc a; invc2 b; color a = Black\ \ invc (baldR a x b)" by (induct a x b rule: baldR.induct) (simp_all add: invc_baliL mkNode_def) lemma invc2_baldR: "\ invc l; invc2 r \ \invc2 (baldR l x r)" by (induct l x r rule: baldR.induct) (auto simp: invc_baliL paint_invc2 invc2I mkNode_def) lemma invh_combine: "\ invh l; invh r; bheight l = bheight r \ \ invh (combine l r) \ bheight (combine l r) = bheight l" by (induct l r rule: combine.induct) (auto simp: invh_baldL_Black split: tree.splits tcolor.splits) lemma invc_combine: assumes "invc l" "invc r" shows "color l = Black \ color r = Black \ invc (combine l r)" "invc2 (combine l r)" using assms by (induct l r rule: combine.induct) (auto simp: invc_baldL invc2I mkNode_def split: tree.splits tcolor.splits) lemma neq_LeafD: "t \ Leaf \ \l x c r. t = Node l (x,c) r" by(cases t) auto lemma del_invc_invh: "invh t \ invc t \ invh (del x t) \ (color t = Red \ bheight (del x t) = bheight t \ invc (del x t) \ color t = Black \ bheight (del x t) = bheight t - 1 \ invc2 (del x t))" proof (induct x t rule: del.induct) case (2 x _ y _ c) have "x = y \ x < y \ x > y" by auto thus ?case proof (elim disjE) assume "x = y" with 2 show ?thesis by (cases c) (simp_all add: invh_combine invc_combine) next assume "x < y" with 2 show ?thesis by(cases c) (auto simp: invh_baldL_invc invc_baldL invc2_baldL mkNode_def dest: neq_LeafD) next assume "y < x" with 2 show ?thesis by(cases c) (auto simp: invh_baldR_invc invc_baldR invc2_baldR mkNode_def dest: neq_LeafD) qed qed auto lemma invpst_baldR: "invpst l \ invpst r \ invpst (baldR l a r)" by (cases "(l,a,r)" rule: baldR.cases) (auto simp: invpst_baliL) lemma invpst_baldL: "invpst l \ invpst r \ invpst (baldL l a r)" by (cases "(l,a,r)" rule: baldL.cases) (auto simp: invpst_baliR) lemma invpst_combine: "invpst l \ invpst r \ invpst (combine l r)" by(induction l r rule: combine.induct) (auto split: tree.splits tcolor.splits simp: invpst_baldR invpst_baldL) lemma invpst_del: "invpst t \ invpst (del x t)" by(induct x t rule: del.induct) (auto simp: invpst_baldR invpst_baldL invpst_combine) theorem rbt_delete: "rbt t \ rbt (delete k t)" apply (clarsimp simp: delete_def rbt_def) apply (frule (1) del_invc_invh[where x=k]) apply (auto simp: invc_paint_Black invh_paint color_paint_Black invpst_del) done lemma rbt_getmin_ismin: - "rbt t \ t\Leaf \ is_min2 (pst_getmin t) (set (inorder t))" + "rbt t \ t\Leaf \ is_min2 (pst_getmin t) (set_tree t)" unfolding rbt_def by (simp add: pst_getmin_ismin) definition "rbt_is_empty t \ t = Leaf" lemma rbt_is_empty: "rbt_is_empty t \ inorder t = []" by (cases t) (auto simp: rbt_is_empty_def) definition empty where "empty = Leaf" subsection \Overall Correctness\ interpretation PM: PrioMap_by_Ordered where empty = empty and lookup = lookup and update = update and delete = delete and inorder = inorder and inv = "rbt" and is_empty = rbt_is_empty and getmin = pst_getmin apply standard apply (auto simp: lookup_map_of inorder_update inorder_delete rbt_update rbt_delete rbt_Leaf rbt_is_empty empty_def dest: rbt_getmin_ismin) done end diff --git a/thys/Simpl/hoare.ML b/thys/Simpl/hoare.ML --- a/thys/Simpl/hoare.ML +++ b/thys/Simpl/hoare.ML @@ -1,3409 +1,3410 @@ (* Title: hoare.ML Author: Norbert Schirmer, TU Muenchen Copyright (C) 2004-2007 Norbert Schirmer This library is free software; you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation; either version 2.1 of the License, or (at your option) any later version. This library is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with this library; if not, write to the Free Software Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 USA *) signature HOARE = sig datatype hoareMode = Partial | Total val gen_proc_rec: Proof.context -> hoareMode -> int -> thm datatype state_kind = Record | Function datatype par_kind = In | Out val deco: string val proc_deco: string val par_deco: string -> string val chopsfx: string -> string -> string val is_state_var: string -> bool val extern: Proof.context -> string -> string val remdeco: Proof.context -> string -> string val remdeco': string -> string val undeco: Proof.context -> term -> term val varname: string -> string val resuffix: string -> string -> string -> string type proc_info = {params: ((par_kind * string) list), recursive: bool, state_kind: state_kind} type hoare_tac = (bool -> int -> tactic) -> Proof.context -> hoareMode -> int -> tactic type hoare_data = {proc_info: proc_info Symtab.table, active_procs: string list list, default_state_kind: state_kind, generate_guard: (stamp * (Proof.context -> term -> term option)), wp_tacs: (string * hoare_tac) list, hoare_tacs: (string * hoare_tac) list, vcg_simps: thm list} val get_data: Proof.context -> hoare_data val get_params: string -> Proof.context -> (par_kind * string) list option val get_default_state_kind: Proof.context -> state_kind val get_state_kind: string -> Proof.context -> state_kind option val clique_name: string list -> string val install_generate_guard: (Proof.context -> term -> term option) -> Context.generic -> Context.generic val generate_guard: Proof.context -> term -> term option val BasicSimpTac: Proof.context -> state_kind -> bool -> thm list -> (int -> tactic) -> int -> tactic val hoare: (Proof.context -> Proof.method) context_parser val hoare_raw: (Proof.context -> Proof.method) context_parser val vcg: (Proof.context -> Proof.method) context_parser val vcg_step: (Proof.context -> Proof.method) context_parser val hoare_rule: (Proof.context -> Proof.method) context_parser val add_foldcongsimps: thm list -> theory -> theory val get_foldcong_ss : theory -> simpset val add_foldcongs : thm list -> theory -> theory val modeqN : string val modexN : string val implementationN : string val specL : string val vcg_tac : string -> string -> string list -> Proof.context -> int -> tactic val hoare_rule_tac : Proof.context -> thm list -> int -> tactic datatype 'a bodykind = BodyTyp of 'a | BodyTerm of 'a val proc_specs : (bstring * string) list parser val add_params : morphism -> string -> (par_kind * string) list -> Context.generic -> Context.generic val set_default_state_kind : state_kind -> Context.generic -> Context.generic val add_state_kind : morphism -> string -> state_kind -> Context.generic -> Context.generic val add_recursive : morphism -> string -> Context.generic -> Context.generic end; structure Hoare: HOARE = struct (* Misc *) val record_vanish = Attrib.setup_config_bool @{binding hoare_record_vanish} (K true); val use_generalise = Attrib.setup_config_bool @{binding hoare_use_generalise} (K false); val sort_variables = Attrib.setup_config_bool @{binding hoare_sort_variables} (K true); val use_cond_inv_modifies = Attrib.setup_config_bool @{binding hoare_use_cond_inv_modifies} (K true); val hoare_trace = Attrib.setup_config_bool @{binding hoare_trace} (K false); val body_def_sfx = "_body"; val programN = "\"; val hoare_ctxtL = "hoare"; val specL = "_spec"; val procL = "_proc"; val bodyP = "_impl"; val modifysfx = "_modifies"; val modexN = "Hoare.mex"; val modeqN = "Hoare.meq"; val KNF = @{const_name StateFun.K_statefun}; (* Some abstract syntax operations *) val Trueprop = HOLogic.mk_Trueprop; infix 0 ===; val (op ===) = Trueprop o HOLogic.mk_eq; fun is_empty_set (Const (@{const_name Orderings.bot}, _)) = true | is_empty_set _ = false; fun mk_Int Ts A B = let val T = fastype_of1 (Ts, A) in Const (@{const_name Lattices.inf}, T --> T --> T) $ A $ B end; fun mk_Un T (A, B) = Const (@{const_name Lattices.sup}, T --> T --> T) $ A $ B; fun dest_Un (Const (@{const_name Lattices.sup}, _) $ t1 $ t2) = dest_Un t1 @ dest_Un t2 | dest_Un t = [t] fun mk_UN' dT rT t = let val dTs = HOLogic.mk_setT dT; val rTs = HOLogic.mk_setT rT; in Const (@{const_name Complete_Lattices.Sup}, rTs --> rT) $ (Const (@{const_name image}, (dT --> rT) --> dTs --> rTs) $ t $ Const (@{const_name Orderings.top}, dTs)) end; fun mk_UN ((x, T), P) = mk_UN' T (fastype_of P) (absfree (x, T) P); fun dest_UN (Const (@{const_name Complete_Lattices.Sup}, _) $ (Const (@{const_name Set.image}, _) $ Abs (x, T, t) $ Const (@{const_name Orderings.top}, _))) = let val (vars, body) = dest_UN t in ((x, T) :: vars, body) end | dest_UN t = ([], t); fun tap_UN (Const (@{const_name Complete_Lattices.Sup}, _) $ (Const (@{const_name Set.image}, _) $ t $ Const (@{const_name Orderings.top}, _))) = SOME t | tap_UN _ = NONE; (* Fetching the rules *) datatype hoareMode = Partial | Total fun get_rule p t Partial = p | get_rule p t Total = t val Guard = get_rule @{thm HoarePartial.Guard} @{thm HoareTotal.Guard}; val GuardStrip = get_rule @{thm HoarePartial.GuardStrip} @{thm HoareTotal.GuardStrip}; val GuaranteeAsGuard = get_rule @{thm HoarePartial.GuaranteeAsGuard} @{thm HoareTotal.GuaranteeAsGuard}; val Guarantee = get_rule @{thm HoarePartial.Guarantee} @{thm HoareTotal.Guarantee}; val GuaranteeStrip = get_rule @{thm HoarePartial.GuaranteeStrip} @{thm HoareTotal.GuaranteeStrip}; val GuardsNil = get_rule @{thm HoarePartial.GuardsNil} @{thm HoareTotal.GuardsNil}; val GuardsCons = get_rule @{thm HoarePartial.GuardsCons} @{thm HoareTotal.GuardsCons}; val GuardsConsGuaranteeStrip = get_rule @{thm HoarePartial.GuardsConsGuaranteeStrip} @{thm HoareTotal.GuardsConsGuaranteeStrip}; val Skip = get_rule @{thm HoarePartial.Skip} @{thm HoareTotal.Skip}; val Basic = get_rule @{thm HoarePartial.Basic} @{thm HoareTotal.Basic}; val BasicCond = get_rule @{thm HoarePartial.BasicCond} @{thm HoareTotal.BasicCond}; val Spec = get_rule @{thm HoarePartial.Spec} @{thm HoareTotal.Spec}; val SpecIf = get_rule @{thm HoarePartial.SpecIf} @{thm HoareTotal.SpecIf}; val Throw = get_rule @{thm HoarePartial.Throw} @{thm HoareTotal.Throw}; val Raise = get_rule @{thm HoarePartial.raise} @{thm HoareTotal.raise}; val Catch = get_rule @{thm HoarePartial.Catch} @{thm HoareTotal.Catch}; val CondCatch = get_rule @{thm HoarePartial.condCatch} @{thm HoareTotal.condCatch}; val CatchSwap = get_rule @{thm HoarePartial.CatchSwap} @{thm HoareTotal.CatchSwap}; val CondCatchSwap = get_rule @{thm HoarePartial.condCatchSwap} @{thm HoareTotal.condCatchSwap}; val Seq = get_rule @{thm HoarePartial.Seq} @{thm HoareTotal.Seq}; val SeqSwap = get_rule @{thm HoarePartial.SeqSwap} @{thm HoareTotal.SeqSwap}; val BSeq = get_rule @{thm HoarePartial.BSeq} @{thm HoareTotal.BSeq}; val Cond = get_rule @{thm HoarePartial.Cond} @{thm HoareTotal.Cond}; val CondInv'Partial = @{thm HoarePartial.CondInv'}; val CondInv'Total = @{thm HoareTotal.CondInv'}; val CondInv' = get_rule CondInv'Partial CondInv'Total; val SwitchNil = get_rule @{thm HoarePartial.switchNil} @{thm HoareTotal.switchNil}; val SwitchCons = get_rule @{thm HoarePartial.switchCons} @{thm HoareTotal.switchCons}; val CondSwap = get_rule @{thm HoarePartial.CondSwap} @{thm HoareTotal.CondSwap}; val While = get_rule @{thm HoarePartial.While} @{thm HoareTotal.While}; val WhileAnnoG = get_rule @{thm HoarePartial.WhileAnnoG} @{thm HoareTotal.WhileAnnoG}; val WhileAnnoFix = get_rule @{thm HoarePartial.WhileAnnoFix'} @{thm HoareTotal.WhileAnnoFix'}; val WhileAnnoGFix = get_rule @{thm HoarePartial.WhileAnnoGFix} @{thm HoareTotal.WhileAnnoGFix}; val BindR = get_rule @{thm HoarePartial.Bind} @{thm HoareTotal.Bind}; val Block = get_rule @{thm HoarePartial.Block} @{thm HoareTotal.Block}; val BlockSwap = get_rule @{thm HoarePartial.BlockSwap} @{thm HoareTotal.BlockSwap}; val Proc = get_rule @{thm HoarePartial.ProcSpec} @{thm HoareTotal.ProcSpec}; val ProcNoAbr = get_rule @{thm HoarePartial.ProcSpecNoAbrupt} @{thm HoareTotal.ProcSpecNoAbrupt}; val ProcBody = get_rule @{thm HoarePartial.ProcBody} @{thm HoareTotal.ProcBody}; val CallBody = get_rule @{thm HoarePartial.CallBody} @{thm HoareTotal.CallBody}; val FCall = get_rule @{thm HoarePartial.FCall} @{thm HoareTotal.FCall}; val ProcRecSpecs = get_rule @{thm HoarePartial.ProcRecSpecs} @{thm HoareTotal.ProcRecSpecs}; val ProcModifyReturnSameFaults = get_rule @{thm HoarePartial.ProcModifyReturnSameFaults} @{thm HoareTotal.ProcModifyReturnSameFaults}; val ProcModifyReturn = get_rule @{thm HoarePartial.ProcModifyReturn} @{thm HoareTotal.ProcModifyReturn}; val ProcModifyReturnNoAbr = get_rule @{thm HoarePartial.ProcModifyReturnNoAbr} @{thm HoareTotal.ProcModifyReturnNoAbr}; val ProcModifyReturnNoAbrSameFaultsPartial = @{thm HoarePartial.ProcModifyReturnNoAbrSameFaults}; val ProcModifyReturnNoAbrSameFaultsTotal = @{thm HoareTotal.ProcModifyReturnNoAbrSameFaults}; val ProcModifyReturnNoAbrSameFaults = get_rule ProcModifyReturnNoAbrSameFaultsPartial ProcModifyReturnNoAbrSameFaultsTotal; val TrivPost = get_rule @{thm HoarePartial.TrivPost} @{thm HoareTotal.TrivPost}; val TrivPostNoAbr = get_rule @{thm HoarePartial.TrivPostNoAbr} @{thm HoareTotal.TrivPostNoAbr}; val DynProcProcPar = get_rule @{thm HoarePartial.DynProcProcPar} @{thm HoareTotal.DynProcProcPar}; val DynProcProcParNoAbr = get_rule @{thm HoarePartial.DynProcProcParNoAbrupt} @{thm HoareTotal.DynProcProcParNoAbrupt}; val ProcProcParModifyReturn = get_rule @{thm HoarePartial.ProcProcParModifyReturn} @{thm HoareTotal.ProcProcParModifyReturn}; val ProcProcParModifyReturnSameFaultsPartial = @{thm HoarePartial.ProcProcParModifyReturnSameFaults}; val ProcProcParModifyReturnSameFaultsTotal = @{thm HoareTotal.ProcProcParModifyReturnSameFaults}; val ProcProcParModifyReturnSameFaults = get_rule ProcProcParModifyReturnSameFaultsPartial ProcProcParModifyReturnSameFaultsTotal; val ProcProcParModifyReturnNoAbr = get_rule @{thm HoarePartial.ProcProcParModifyReturnNoAbr} @{thm HoareTotal.ProcProcParModifyReturnNoAbr}; val ProcProcParModifyReturnNoAbrSameFaultsPartial = @{thm HoarePartial.ProcProcParModifyReturnNoAbrSameFaults}; val ProcProcParModifyReturnNoAbrSameFaultsTotal = @{thm HoareTotal.ProcProcParModifyReturnNoAbrSameFaults}; val ProcProcParModifyReturnNoAbrSameFaults = get_rule ProcProcParModifyReturnNoAbrSameFaultsPartial ProcProcParModifyReturnNoAbrSameFaultsTotal; val DynCom = get_rule @{thm HoarePartial.DynComConseq} @{thm HoareTotal.DynComConseq}; val AugmentContext = get_rule @{thm HoarePartial.augment_context'} @{thm HoareTotal.augment_context'}; val AugmentEmptyFaults = get_rule @{thm HoarePartial.augment_emptyFaults} @{thm HoareTotal.augment_emptyFaults}; val AsmUN = get_rule @{thm HoarePartial.AsmUN} @{thm HoareTotal.AsmUN}; val SpecAnno = get_rule @{thm HoarePartial.SpecAnno'} @{thm HoareTotal.SpecAnno'}; val SpecAnnoNoAbrupt = get_rule @{thm HoarePartial.SpecAnnoNoAbrupt} @{thm HoareTotal.SpecAnnoNoAbrupt}; val LemAnno = get_rule @{thm HoarePartial.LemAnno} @{thm HoareTotal.LemAnno}; val LemAnnoNoAbrupt = get_rule @{thm HoarePartial.LemAnnoNoAbrupt} @{thm HoareTotal.LemAnnoNoAbrupt}; val singleton_conv_sym = @{thm Set.singleton_conv2} RS sym; val anno_defs = [@{thm Language.whileAnno_def},@{thm Language.whileAnnoG_def},@{thm Language.specAnno_def}, @{thm Language.whileAnnoGFix_def},@{thm Language.whileAnnoFix_def},@{thm Language.lem_def}]; val strip_simps = @{thm Language.strip_simp} :: @{thm Option.option.map(2)} :: @{thms Language.strip_guards_simps}; val normalize_simps = [@{thm Language.while_def}, @{thm Language.bseq_def}, @{thm List.append_Nil}, @{thm List.append_Cons}] @ @{thms List.list.cases} @ @{thms Language.flatten_simps} @ @{thms Language.sequence.simps} @ @{thms Language.normalize_simps} @ @{thms Language.guards.simps} @ [@{thm fst_conv}, @{thm snd_conv}]; val K_rec_convs = []; val K_fun_convs = [@{thm StateFun.K_statefun_apply}, @{thm StateFun.K_statefun_comp}]; val K_convs = K_rec_convs @ K_fun_convs; val K_rec_congs = []; val K_fun_congs = [@{thm StateFun.K_statefun_cong}]; val K_congs = K_rec_congs @ K_fun_congs; (* misc. aux. functions *) (* first_subterm * yields result x of P for first subterm for which P is (SOME x), and all bound variables on the path * to that term *) fun first_subterm_dest P = let fun first abs_vars t = (case P t of SOME x => SOME (abs_vars,x) |_=> (case t of u $ v => (case first abs_vars u of NONE => first abs_vars v | SOME x => SOME x) | Abs (c,T,u) => first (abs_vars @ [(c,T)]) u | _ => NONE)) in first [] end; (* first_subterm * yields first subterm for which P holds, and all bound variables on the path * to that term *) fun first_subterm P = let fun P' t = if P t then SOME t else NONE; in first_subterm_dest P' end; (* max_subterm_dest * yields results of P for all maximal subterms for which P is (SOME x), * and all bound variables on the path to that subterm *) fun max_subterms_dest P = let fun collect abs_vars t = (case P t of SOME x => [(abs_vars,x)] |_=> (case t of u $ v => collect abs_vars u @ collect abs_vars v | Abs (c,T,u) => collect (abs_vars @ [(c,T)]) u | _ => [])) in collect [] end; fun last [] = raise Empty | last [x] = x | last (_::xs) = last xs; fun dest_splits (Const (@{const_name case_prod},_)$Abs (n,T,t)) = (n,T)::dest_splits t | dest_splits (Const (@{const_name case_prod},_)$Abs (n,T,t)$_) = (n,T)::dest_splits t | dest_splits (Abs (n,T,_)) = [(n,T)] | dest_splits _ = []; fun idx eq [] x = ~1 | idx eq (x::rs) y = if eq x y then 0 else let val i = idx eq rs y in if i < 0 then i else i+1 end; fun resuffix sfx1 sfx2 s = suffix sfx2 (unsuffix sfx1 s) handle Fail _ => s; (* state space representation dependent functions *) datatype state_kind = Record | Function fun state_simprocs Record = [Record.simproc] | state_simprocs Function = [Record.simproc, StateFun.lookup_simproc]; fun state_upd_simproc Record = Record.upd_simproc | state_upd_simproc Function = StateFun.update_simproc; fun state_ex_sel_eq_simproc Record = Record.ex_sel_eq_simproc | state_ex_sel_eq_simproc Function = StateFun.ex_lookup_eq_simproc; val state_split_simp_tac = Record.split_simp_tac val state_hierarchy = Record.dest_recTs fun stateT_id T = case (state_hierarchy T) of [] => NONE | Ts => SOME (last Ts); fun globalsT (Type (_, T :: _)) = SOME T | globalsT _ = NONE; fun stateT_ids T = (case stateT_id T of NONE => NONE | SOME sT => (case globalsT T of NONE => SOME [sT] | SOME gT => (case stateT_id gT of NONE => SOME [sT] | SOME gT' => SOME [sT,gT']))); datatype par_kind = In | Out (*** utilities ***) (* utils for variable name decorations *) val deco = "_'"; val proc_deco = "_'proc"; fun par_deco name = deco ^ name ^ deco; fun chopsfx sfx str = (case try (unsuffix sfx) str of SOME s => s | NONE => str) val is_state_var = can (unsuffix deco); (* removes the suffix of the string beginning with deco. * "xys_'a" --> "xys"; * The a is also chopped, since sometimes the bound variables * are renamed, I think SELECT_GOAL in rename_goal is to blame *) fun remdeco' str = let fun chop (p::ps) (x::xs) = chop ps xs | chop [] xs = [] | chop (p::ps) [] = error "remdeco: code should never be reached"; fun remove prf (s as (x::xs)) = if is_prefix (op =) prf s then chop prf s else (x::remove prf xs) | remove prf [] = []; in String.implode (remove (String.explode deco) (String.explode str)) end; fun extern ctxt s = (case try (Proof_Context.extern_const ctxt o Lexicon.unmark_const) s of NONE => s | SOME s' => s'); fun remdeco ctxt s = remdeco' (extern ctxt s); fun undeco ctxt (Const (c, T)) = Const (remdeco ctxt c, T) | undeco ctxt ((f as Const (@{syntax_const "_free"},_)) $ Free (x, T)) = (*f$*)Const (remdeco' x, T) | undeco ctxt (Const _ $ _ $ ((Const (@{syntax_const "_free"},_)) $ Free (x, T))) = (*f$*)Const (remdeco' x, T) | undeco ctxt (Free (c, T)) = Const (remdeco' c, T) | undeco ctxt x = x fun varname x = x ^ deco val dest_string = map (chr o HOLogic.dest_char) o HOLogic.dest_list; fun dest_string' t = (case try dest_string t of SOME s => implode s | NONE => (case t of Free (s,_) => s | Const (s,_) => Long_Name.base_name s | _ => raise TERM ("dest_string'",[t]))) fun is_state_space_var Tids t = let fun is_stateT T = (case stateT_id T of NONE => 0 | SOME id => if member (op =) Tids id then ~1 else 0); in (case t of Const _ $ Abs (_,T,_) => is_stateT T | Free (_,T) => is_stateT T | _ => 0) end; datatype callMode = Static | Parameter fun proc_name Static (Const (p,_)$_) = resuffix deco proc_deco (Long_Name.base_name p) | proc_name Static (Const (@{const_name StateFun.lookup},_)$_$Free (p,_)$_) = suffix proc_deco (remdeco' (Long_Name.base_name p)) | proc_name Static p = dest_string' p | proc_name Parameter (Const (p,_)) = resuffix deco proc_deco (Long_Name.base_name p) | proc_name Parameter (Abs (_,_,Const (p,_)$Bound 0)) = resuffix deco proc_deco (Long_Name.base_name p) | proc_name Parameter (Abs (_,_,Const (@{const_name StateFun.lookup},_)$_$Free (p,_)$_)) = suffix proc_deco (remdeco' (Long_Name.base_name p)) | proc_name _ t = raise TERM ("proc_name",[t]); fun dest_call (Const (@{const_name Language.call},_)$init$pname$return$c) = (init,pname,return,c,Static,true) | dest_call (Const (@{const_name Language.fcall},_)$init$pname$return$_$c) = (init,pname,return,c,Static,true) | dest_call (Const (@{const_name Language.com.Call},_)$pname) = (Bound 0,pname,Bound 0,Bound 0,Static,false) | dest_call (Const (@{const_name Language.dynCall},_)$init$pname$return$c) = (init,pname,return,c,Parameter,true) | dest_call t = raise TERM ("Hoare.dest_call: unexpected term",[t]); fun dest_whileAnno (Const (@{const_name Language.whileAnnoG},_) $gs$b$I$V$c) = (SOME gs,b,I,V,c,false) | dest_whileAnno (Const (@{const_name Language.whileAnno},_) $b$I$V$c) = (NONE,b,I,V,c,false) | dest_whileAnno (Const (@{const_name Language.whileAnnoGFix},_)$gs$b$I$V$c) = (SOME gs,b,I,V,c,true) | dest_whileAnno (Const (@{const_name Language.whileAnnoFix},_) $b$I$V$c) = (NONE,b,I,V,c,true) | dest_whileAnno t = raise TERM ("Hoare.dest_while: unexpected term",[t]); fun dest_Guard (Const (@{const_name Language.com.Guard},_)$f$g$c) = (f,g,c,false) | dest_Guard (Const (@{const_name Language.guaranteeStrip},_)$f$g$c) = (f,g,c,true) | dest_Guard t = raise TERM ("Hoare.dest_guard: unexpected term",[t]); (*** extend theory by procedure definition ***) fun add_declaration name decl thy = thy |> Named_Target.init name |> Local_Theory.declaration {syntax = false, pervasive = false} decl |> Local_Theory.exit |> Proof_Context.theory_of; (* data kind 'HOL/hoare' *) type proc_info = {params: ((par_kind * string) list), recursive: bool, state_kind: state_kind} type hoare_tac = (bool -> int -> tactic) -> Proof.context -> hoareMode -> int -> tactic; type hoare_data = {proc_info: proc_info Symtab.table, active_procs: string list list, default_state_kind: state_kind, generate_guard: (stamp * (Proof.context -> term -> term option)), wp_tacs: (string * hoare_tac) list, hoare_tacs: (string * hoare_tac) list, vcg_simps: thm list}; fun make_hoare_data proc_info active_procs default_state_kind generate_guard wp_tacs hoare_tacs vcg_simps = {proc_info = proc_info, active_procs = active_procs, default_state_kind = default_state_kind, generate_guard = generate_guard, wp_tacs = wp_tacs, hoare_tacs = hoare_tacs, vcg_simps = vcg_simps}; structure Hoare_Data = Generic_Data ( type T = hoare_data; val empty = make_hoare_data (Symtab.empty: proc_info Symtab.table) ([]:string list list) (Function) (stamp (),(K (K NONE)): Proof.context -> term -> term option) ([]:(string * hoare_tac) list) ([]:(string * hoare_tac) list) ([]:thm list); val extend = I; (* FIXME exponential blowup due to append !? *) fun merge ({proc_info = proc_info1, active_procs = active_procs1, default_state_kind = _, generate_guard = (stmp1,generate_gaurd1), wp_tacs = wp_tacs1, hoare_tacs = hoare_tacs1, vcg_simps = vcg_simps1}, {proc_info = proc_info2, active_procs = active_procs2, default_state_kind = default_state_kind2, generate_guard = (stmp2, _), wp_tacs = wp_tacs2, hoare_tacs = hoare_tacs2, vcg_simps=vcg_simps2}) : T = if stmp1=stmp2 then make_hoare_data (Symtab.merge (K true) (proc_info1,proc_info2)) (active_procs1 @ active_procs2) (default_state_kind2) (stmp1,generate_gaurd1) (wp_tacs1 @ wp_tacs2) (hoare_tacs1 @ hoare_tacs2) (Thm.merge_thms (vcg_simps1,vcg_simps2)) else error ("Theories have different aux. functions to generate guards") ); val get_data = Hoare_Data.get o Context.Proof; (* access 'params' *) fun mk_free ctxt name = let val ctxt' = Context.proof_of ctxt; val n' = Variable.intern_fixed ctxt' name |> perhaps Long_Name.dest_hidden; val T' = Proof_Context.infer_type ctxt' (n', dummyT) handle ERROR _ => dummyT in (Free (n',T')) end; fun morph_name ctxt phi name = (case Morphism.term phi (mk_free ctxt name) of Free (x,_) => x | _ => name); datatype 'a bodykind = BodyTyp of 'a | BodyTerm of 'a fun set_default_state_kind sk ctxt = let val {proc_info,active_procs,default_state_kind,generate_guard,wp_tacs,hoare_tacs, vcg_simps,...} = Hoare_Data.get ctxt; val data = make_hoare_data proc_info active_procs sk generate_guard wp_tacs hoare_tacs vcg_simps; in Hoare_Data.put data ctxt end; val get_default_state_kind = #default_state_kind o get_data; fun add_active_procs phi ps ctxt = let val {proc_info,active_procs,default_state_kind,generate_guard,wp_tacs,hoare_tacs, vcg_simps,...} = Hoare_Data.get ctxt; val data = make_hoare_data proc_info ((map (morph_name ctxt phi) ps)::active_procs) default_state_kind generate_guard wp_tacs hoare_tacs vcg_simps; in Hoare_Data.put data ctxt end; fun add_hoare_tacs tacs ctxt = let val {proc_info,active_procs, default_state_kind, generate_guard,wp_tacs,hoare_tacs, vcg_simps,...} = Hoare_Data.get ctxt; val data = make_hoare_data proc_info active_procs default_state_kind generate_guard wp_tacs (hoare_tacs@tacs) vcg_simps; in Hoare_Data.put data ctxt end; fun map_vcg_simps f ctxt = let val {proc_info,active_procs,default_state_kind,generate_guard,wp_tacs,hoare_tacs, vcg_simps,...} = Hoare_Data.get ctxt; val data = make_hoare_data proc_info active_procs default_state_kind generate_guard wp_tacs hoare_tacs (f vcg_simps); in Hoare_Data.put data ctxt end; fun thy_attrib f = Thm.declaration_attribute (fn thm => map_vcg_simps (f thm)); val vcg_simpadd = Thm.add_thm val vcg_simpdel = Thm.del_thm val vcg_simp_add = thy_attrib vcg_simpadd; val vcg_simp_del = thy_attrib vcg_simpdel; (* add 'procedure' *) fun mk_proc_info params recursive state_kind = {params=params,recursive=recursive,state_kind=state_kind}; val empty_proc_info = {params=[],recursive=false,state_kind=Record}; fun map_proc_info_params f {params,recursive,state_kind} = mk_proc_info (f params) recursive state_kind; fun map_proc_info_recursive f {params,recursive,state_kind} = mk_proc_info params (f recursive) state_kind; fun map_proc_info_state_kind f {params,recursive,state_kind} = mk_proc_info params recursive (f state_kind); fun add_params phi name frmls ctxt = let val {proc_info,active_procs,default_state_kind,generate_guard,wp_tacs,hoare_tacs, vcg_simps,...} = Hoare_Data.get ctxt; val params = map (apsnd (morph_name ctxt phi)) frmls; val f = map_proc_info_params (K params); val default = f empty_proc_info; val proc_info' = Symtab.map_default (morph_name ctxt phi name,default) f proc_info; val data = make_hoare_data proc_info' active_procs default_state_kind generate_guard wp_tacs hoare_tacs vcg_simps; in Hoare_Data.put data ctxt end; fun get_params name ctxt = Option.map #params (Symtab.lookup (#proc_info (get_data ctxt)) name); fun add_recursive phi name ctxt = let val {proc_info,active_procs,default_state_kind,generate_guard,wp_tacs,hoare_tacs, vcg_simps,...} = Hoare_Data.get ctxt; val f = map_proc_info_recursive (K true); val default = f empty_proc_info; val proc_info'= Symtab.map_default (morph_name ctxt phi name,default) f proc_info; val data = make_hoare_data proc_info' active_procs default_state_kind generate_guard wp_tacs hoare_tacs vcg_simps; in Hoare_Data.put data ctxt end; fun get_recursive name ctxt = Option.map #recursive (Symtab.lookup (#proc_info (get_data ctxt)) name); fun add_state_kind phi name sk ctxt = let val {proc_info,active_procs,default_state_kind,generate_guard,wp_tacs,hoare_tacs, vcg_simps,...} = Hoare_Data.get ctxt; val f = map_proc_info_state_kind (K sk); val default = f empty_proc_info; val proc_info'= Symtab.map_default (morph_name ctxt phi name,default) f proc_info; val data = make_hoare_data proc_info' active_procs default_state_kind generate_guard wp_tacs hoare_tacs vcg_simps; in Hoare_Data.put data ctxt end; fun get_state_kind name ctxt = Option.map #state_kind (Symtab.lookup (#proc_info (get_data ctxt)) name); fun install_generate_guard f ctxt = let val {proc_info,active_procs, default_state_kind, generate_guard,wp_tacs,hoare_tacs, vcg_simps,...} = Hoare_Data.get ctxt; val data = make_hoare_data proc_info active_procs default_state_kind (stamp (), f) wp_tacs hoare_tacs vcg_simps in Hoare_Data.put data ctxt end; fun generate_guard ctxt = snd (#generate_guard (get_data ctxt)) ctxt; fun check_procedures_definition procs thy = let val ctxt = Proof_Context.init_global thy; fun already_defined name = if is_some (get_params name ctxt) then ["procedure " ^ quote name ^ " already defined"] else [] val err_already_defined = maps (already_defined o #1) procs; fun duplicate_procs names = (case duplicates (op =) names of [] => [] | dups => ["Duplicate procedures " ^ commas_quote dups]); val err_duplicate_procs = duplicate_procs (map #1 procs); fun duplicate_pars name pars = (case duplicates (op =) (map fst pars) of [] => [] | dups => ["Duplicate parameters in procedure " ^ quote name ^ ": " ^ commas_quote dups]); val err_duplicate_pars = maps (fn (name,inpars,outpars,locals,_,_,_) => duplicate_pars name (inpars @ locals) @ duplicate_pars name (outpars @ locals)) procs; (* FIXME: Check that no global variables are used as result parameters *) val errs = err_already_defined @ err_duplicate_procs @ err_duplicate_pars; in if null errs then () else error (cat_lines errs) end; fun add_parameter_info phi cname (name,(inpars,outpars,state_kind)) ctxt = let fun par_deco' T = if T = "" then deco else par_deco (cname name); val pars = map (fn (par,T) => (In,suffix (par_deco' T) par)) inpars@ map (fn (par,T) => (Out,suffix (par_deco' T) par)) outpars; val ctxt_decl = ctxt |> add_params phi name pars |> add_state_kind phi name state_kind in ctxt_decl end; fun mk_loc_exp xs = let fun mk_expr s = (s,(("",false),(Expression.Named [],[]))) in (map mk_expr xs,[]) end; val parametersN = "_parameters"; val variablesN = "_variables"; val signatureN = "_signature"; val bodyN = "_body"; val implementationN = "_impl"; val cliqueN = "_clique"; val clique_namesN = "_clique_names"; val NoBodyN = @{const_name Vcg.NoBody}; val statetypeN = "StateType"; val proc_nameT = HOLogic.stringT; fun expression_no_pos (expr, fixes) : Expression.expression = (map (fn (name, inst) => ((name, Position.none), inst)) expr, fixes); fun add_locale name expr elems thy = thy |> Expression.add_locale (Binding.name name) (Binding.name name) expr elems |> snd |> Local_Theory.exit; fun add_locale' name expr elems thy = thy |> Expression.add_locale (Binding.name name) (Binding.name name) expr elems ||> Local_Theory.exit; fun add_locale_cmd name expr elems thy = thy |> Expression.add_locale_cmd (Binding.name name) (Binding.name name) (expression_no_pos expr) elems |> snd |> Local_Theory.exit; fun read_typ thy raw_T env = let val ctxt' = Proof_Context.init_global thy |> fold (Variable.declare_typ o TFree) env; val T = Syntax.read_typ ctxt' raw_T; val env' = Term.add_tfreesT T env; in (T, env') end; fun add_variable_statespaces (cname, (inpars, outpars, locvars)) thy = let val inpars' = if forall (fn (_,T) => T = "") inpars then [] else inpars; val outpars' = if forall (fn (_,T) => T = "") outpars then [] else outpars; fun prep_comp (n, T) env = let val (T', env') = read_typ thy T env handle ERROR msg => cat_error msg ("The error(s) above occurred in component " ^ quote n) in ((n, T'), env') end; val (in_outs,in_out_env) = fold_map prep_comp (distinct (op =) (inpars'@outpars')) []; val (locs,var_env) = fold_map prep_comp locvars in_out_env; val parSP = cname ^ parametersN; val in_outs' = map (apfst (suffix (par_deco cname))) in_outs; val in_out_args = map fst in_out_env; val varSP = cname ^ variablesN; val locs' = map (apfst (suffix (par_deco cname))) locs; val var_args = map fst var_env; in if null inpars' andalso null outpars' andalso null locvars then thy |> add_locale_cmd parSP ([],[]) [] |> Proof_Context.theory_of |> add_locale_cmd varSP ([],[]) [] |> Proof_Context.theory_of else thy |> StateSpace.define_statespace_i (SOME statetypeN) in_out_args parSP [] in_outs' |> StateSpace.define_statespace_i (SOME statetypeN) var_args varSP [((cname, false), ((map TFree in_out_env),parSP,[]))] locs' end; fun intern_locale thy = Locale.intern thy #> perhaps Long_Name.dest_hidden; fun apply_in_context thy lexp f t = let fun name_variant lname = if intern_locale thy lname = lname then lname else name_variant (lname ^ "'"); in thy (* Create a dummy locale in dummy theory just to read the term *) |> add_locale_cmd (name_variant "foo") lexp [] |> (fn ctxt => f ctxt t) end; fun add_abbrev loc mode name spec thy = thy |> Named_Target.init loc |> (fn lthy => let val t = Syntax.read_term (Local_Theory.target_of lthy) spec; in Local_Theory.abbrev mode ((Binding.name name, NoSyn), t) lthy end) |> #2 |> Local_Theory.exit |> Proof_Context.theory_of; exception TOPSORT of string fun topsort less [] = [] | topsort less xs = let fun list_all P xs = fold (fn x => fn b => b andalso P x) xs true; fun split_min n (x::xs) = if n=0 then raise TOPSORT "no minimum in list" else if list_all (less x) xs then (x,xs) else split_min (n-1) (xs@[x]); fun tsort [] = [] | tsort xs = let val (x,xs') = split_min (length xs) xs; in x::tsort xs' end; in tsort xs end; fun clique_name clique = (foldr1 (fn (a,b) => a ^ "_" ^ b) (map (unsuffix proc_deco) clique)); fun error_to_warning msg f thy = f thy handle ERROR msg' => (warning (msg' ^ "\n" ^ msg); thy); fun procedures_definition locname procs thy = let val procs' = map (fn (name,a,b,c,d,e,f) => (suffix proc_deco name,a,b,c,d,e,f)) procs; val _ = check_procedures_definition procs' thy; val name_pars = map (fn (name,inpars,outpars,_,_,_,sk) => (name,(inpars,outpars,sk))) procs'; val name_vars = map (fn (name,inpars,outpars,locals,_,_,_) => (name,(inpars,outpars,locals))) procs'; val name_body = map (fn (name,_,_,_,body,_,_) => (name,body)) procs'; val name_pars_specs = map (fn (name,inpars,outpars,_,_,specs,sk) => (name,(inpars,outpars,sk),specs)) procs'; val names = map #1 procs'; val sk = #7 (hd procs'); val thy = thy |> Context.theory_map (set_default_state_kind sk); val (all_callss,cliques,is_recursive,has_body) = let val ctxt = Context.Theory thy |> fold (add_parameter_info Morphism.identity (unsuffix proc_deco)) name_pars |> StateSpace.set_silent true fun read_body (_, body) = Syntax.read_term (Context.proof_of ctxt) body; val bodies = map read_body name_body; fun dcall t = (case try dest_call t of SOME (_,p,_,_,m,_) => SOME (proc_name m p) | _ => NONE); fun in_names x = if member (op =) names x then SOME x else NONE; fun add_edges n = fold (fn x => Graph.add_edge (n, x)); val all_callss = map (map snd o max_subterms_dest dcall) bodies; val callss = map (map_filter in_names) all_callss; val graph = fold (fn n => Graph.new_node (n, ())) names Graph.empty; val graph' = fold2 add_edges names callss graph; fun idx x = find_index (fn y => x=y) names; fun name_ord (a,b) = int_ord (idx a, idx b); val cliques = Graph.strong_conn graph'; val cliques' = map (sort name_ord) cliques; val my_calls = these o AList.lookup (op =) (names ~~ map (distinct (op =)) callss); val my_body = AList.lookup (op =) (names ~~ bodies); fun is_recursive n = exists (fn [_] => false | ns => member (op =) ns n) (Graph.strong_conn graph'); fun has_body n = (case my_body n of SOME (Const (c,_)) => c <> NoBodyN | _ => true) fun clique_less c1 c2 = null (inter (op =) (distinct (op =) (maps my_calls c1)) c2); val cliques'' = topsort clique_less cliques'; in (all_callss,cliques'',is_recursive,has_body) end; (* cliques may only depend on ones to the left, so it is safe to * add the locales from the left to the right. *) fun my_clique n = Library.find_first (fn c => member (op =) c n) cliques; fun lname sfx clique = suffix sfx (clique_name clique); fun cname n = clique_name (the (my_clique n)); fun parameter_info_decl phi = fold (add_parameter_info phi cname) name_pars; fun get_loc sfx clique n = if member (op =) clique n then NONE else SOME (resuffix proc_deco sfx n); fun parent_locales thy sfx clique = let val calls = distinct (op =) (flat (map_filter (AList.lookup (op =) (names ~~ all_callss)) clique)); in map (intern_locale thy) (distinct (op =) (map_filter (get_loc sfx clique) calls)) end; val names_all_callss = names ~~ map (distinct (op =)) all_callss; val get_calls = the o AList.lookup (op =) names_all_callss; fun clique_vars clique = let fun add name (ins,outs,locs) = let val (nins,nouts,nlocs) = the (AList.lookup (op =) name_vars name) in (ins@nins,outs@nouts,locs@nlocs) end; val (is,os,ls) = fold add clique ([],[],[]); in (lname "" clique, (distinct (op =) is, distinct (op =) os, distinct (op =) ls)) end; fun add_signature_locale (cname, name) thy = let val name' = unsuffix proc_deco name; val fixes = [Element.Fixes [(Binding.name name, SOME proc_nameT, NoSyn)]]; (* FIXME: may use HOLogic.typeT as soon as locale type-inference works properly *) val pE = mk_loc_exp [intern_locale thy (suffix parametersN cname)]; val sN = suffix signatureN name'; in thy |> add_locale sN pE fixes |> Proof_Context.theory_of |> (fn thy => add_declaration (intern_locale thy sN) parameter_info_decl thy) end; fun mk_bdy_def read_term name = let val name' = unsuffix proc_deco name; val bdy = read_term (the (AList.lookup (op =) name_body name)); val bdy_defN = suffix body_def_sfx name'; val b = Binding.name bdy_defN; in ((b, NoSyn), ((Thm.def_binding b, []), bdy)) end; fun add_body_locale (name, _) thy = let val name' = unsuffix proc_deco name; val callees = filter_out (fn n => n = name) (get_calls name) val fixes = [Element.Fixes [(Binding.name name, SOME proc_nameT, NoSyn)]]; (* FIXME: may use HOLogic.typeT as soon as locale type-inference works properly *) val pE = mk_loc_exp (map (intern_locale thy) ([lname variablesN (the (my_clique name))]@ the_list locname@ map (resuffix proc_deco signatureN) callees)); fun def lthy = let val read = Syntax.read_term (Context.proof_map (add_active_procs Morphism.identity (the (my_clique name))) (Local_Theory.target_of lthy)) in mk_bdy_def read name end; fun add_decl_and_def lname ctxt = ctxt |> Proof_Context.theory_of |> Named_Target.init lname |> Local_Theory.declaration {syntax = false, pervasive = false} parameter_info_decl |> (fn lthy => if has_body name then snd (Local_Theory.define (def lthy) lthy) else lthy) |> Local_Theory.exit |> Proof_Context.theory_of; in thy |> add_locale' (suffix bodyN name') pE fixes |-> add_decl_and_def end; - fun mk_def_eq _ read_term name = + fun mk_def_eq thy read_term name = if has_body name then let (* FIXME: All the read_term stuff is just because type-inference/abbrevs for * new locale elements does not work right now; * We read the term to expand the abbreviations, then we print it again * (without folding the abbreviation) and reread as string *) val name' = unsuffix proc_deco name; val bdy_defN = suffix body_def_sfx name'; val rhs = read_term ("Some " ^ bdy_defN); val nt = read_term name; val Free (gamma,_) = read_term programN; val eq = HOLogic.Trueprop$ HOLogic.mk_eq (Free (gamma,fastype_of nt --> fastype_of rhs)$nt,rhs) + val consts = Sign.consts_of thy; val eqs = - YXML.string_of_body (Term_XML.Encode.term (map_types (K dummyT) eq)); + YXML.string_of_body (Term_XML.Encode.term consts (Consts.dummy_types consts eq)); val assms = Element.Assumes [((Binding.name (suffix bodyP name'), []),[(eqs,[])])] in [assms] end else []; fun add_impl_locales clique thy = let val cliqN = lname cliqueN clique; val cnamesN = lname clique_namesN clique; val multiple_procs = length clique > 1; val add_distinct_procs_namespace = if multiple_procs then StateSpace.namespace_definition cnamesN proc_nameT ([],[]) [] clique else I; val bodyLs = map (suffix bodyN o unsuffix proc_deco) clique; fun pE thy = mk_loc_exp (map (intern_locale thy) (hoare_ctxtL::bodyLs) @ (parent_locales thy implementationN clique) @ (if multiple_procs then [intern_locale thy cnamesN] else [])); fun read_term thy = apply_in_context thy (pE thy) Syntax.read_term; fun elems thy = maps (mk_def_eq thy (read_term thy)) clique; fun add_recursive_info phi name = if is_recursive name then (add_recursive phi name) else I; fun proc_declaration phi = add_active_procs phi clique; fun recursive_declaration phi ctxt = ctxt |> fold (add_recursive_info phi) clique; fun add_impl_locale name thy = let val implN = suffix implementationN (unsuffix proc_deco name); val parentN = intern_locale thy cliqN val parent = mk_loc_exp [parentN]; in thy |> add_locale_cmd implN parent [] |> Proof_Context.theory_of |> (fn thy => Interpretation.global_sublocale parentN (mk_loc_exp [intern_locale thy implN]) [] thy) |> Proof.global_terminal_proof ((Method.Basic (fn ctxt => Method.SIMPLE_METHOD (Locale.intro_locales_tac {strict = true, eager = false} ctxt [])), Position.no_range), NONE) |> Proof_Context.theory_of end; in thy |> add_distinct_procs_namespace |> (fn thy => add_locale_cmd cliqN (pE thy) (elems thy) thy) |> Proof_Context.theory_of |> fold add_impl_locale clique |> (fn thy => add_declaration (intern_locale thy cliqN) proc_declaration thy) |> (fn thy => add_declaration (intern_locale thy cliqN) recursive_declaration thy) end; fun add_spec_locales (name, _, specs) thy = let val name' = unsuffix proc_deco name; val ps = (suffix signatureN name' :: the_list locname); val ps' = hoare_ctxtL :: ps ; val pE = mk_loc_exp (map (intern_locale thy) ps) val pE' = mk_loc_exp (map (intern_locale thy) ps') fun read thy = apply_in_context thy (mk_loc_exp [intern_locale thy (suffix cliqueN (cname name))]) (Syntax.read_prop); fun proc_declaration phi = (*parameter_info_decl phi o already in signature *) add_active_procs phi (the (my_clique name)); fun add_locale'' (thm_name,spec) thy = let val spec' = read thy spec; val elem = Element.Assumes [((Binding.name thm_name, []), [(spec', [])])]; in thy |> add_locale thm_name pE' [elem] |> Proof_Context.theory_of |> (fn thy => add_declaration (intern_locale thy thm_name) proc_declaration thy) |> error_to_warning ("abbreviation: '" ^ thm_name ^ "' not added") (add_abbrev (intern_locale thy (suffix cliqueN (cname name))) Syntax.mode_input thm_name spec) end; in thy |> fold add_locale'' specs end; in thy |> fold (add_variable_statespaces o clique_vars) cliques |> fold (fn c => fold (fn n => add_signature_locale (lname "" c, n)) c) cliques |> fold add_body_locale name_pars |> fold add_impl_locales cliques |> fold add_spec_locales name_pars_specs end; (********************* theory extender interface ********************************) (** package setup **) (* outer syntax *) val var_declP = Parse.name -- (@{keyword "::"} |-- Parse.!!! Parse.embedded); val var_declP' = Parse.name >> (fn n => (n,"")); val localsP = Scan.repeat var_declP; val argP = var_declP; val argP' = var_declP'; val not_eqP = Scan.ahead (Scan.unless @{keyword "="} (Scan.one (K true))) val proc_decl_statespace = (Parse.short_ident --| @{keyword "("}) -- ((Parse.list argP) -- (Scan.optional (@{keyword "|"} |-- Parse.list argP) []) --| @{keyword ")"}) --| not_eqP val proc_decl_record = (Parse.short_ident --| @{keyword "("}) -- ((Parse.list argP') -- (Scan.optional (@{keyword "|"} |-- Parse.list argP') []) --| @{keyword ")"}) --| Scan.option @{keyword "="} val proc_decl = proc_decl_statespace >> pair Function || proc_decl_record >> pair Record; val loc_decl = Scan.optional (@{keyword "where"} |-- localsP --| @{keyword "in"}) [] val proc_body = Parse.embedded (*>> BodyTerm*) fun proc_specs x = (Parse.!!! (Scan.repeat (Parse_Spec.thm_name ":" -- Parse.embedded)) >> map (fn ((thm_name, _), prop) => (Binding.name_of thm_name, prop))) x val par_loc = Scan.option (@{keyword "("} |-- @{keyword "imports"} |-- Parse.name --| @{keyword ")"}); val _ = Outer_Syntax.command @{command_keyword procedures} "define procedures" (par_loc -- (Parse.and_list1 (proc_decl -- loc_decl -- proc_body -- proc_specs)) >> (fn (loc,decls) => let val decls' = map (fn ((((state_kind,(name,(ins,outs))),ls),bdy),specs) => (name,ins,outs,ls,bdy,specs,state_kind)) decls in Toplevel.theory (procedures_definition loc decls') end)); val _ = Outer_Syntax.command @{command_keyword hoarestate} "define state space for hoare logic" (StateSpace.statespace_decl >> (fn ((args,name),(parents,comps)) => Toplevel.theory (StateSpace.define_statespace args name parents (map (apfst (suffix deco)) comps)))); (*************************** Auxiliary Functions for integration of ********************) (*************************** automatic program analysers ********************) fun dest_conjs t = (case HOLogic.dest_conj t of [t1,t2] => dest_conjs t1 @ dest_conjs t2 | ts => ts); fun split_guard (Const (@{const_name Collect},CT)$(Abs (s,T,t))) = let fun mkCollect t = Const (@{const_name Collect},CT)$(Abs (s,T,t)); in map mkCollect (dest_conjs t) end | split_guard t = [t]; fun split_guards gs = let fun norm c f g = map (fn g => c$f$g) (split_guard g); fun norm_guard ((c as Const (@{const_name Pair},_))$f$g) = norm c f g | norm_guard ((c as Const (@{const_name Language.guaranteeStripPair},_))$f$g) = norm c f g | norm_guard t = [t]; in maps norm_guard (HOLogic.dest_list gs) end fun fold_com f t = let (* traverse does not descend into abstractions, like in DynCom, call, etc. *) fun traverse cnt (c as Const (@{const_name Language.com.Skip},_)) = (cnt,f cnt c [] []) | traverse cnt ((c as Const (@{const_name Language.com.Basic},_))$g) = (cnt, f cnt c [g] []) | traverse cnt ((c as Const (@{const_name Language.com.Spec},_))$r) = (cnt, f cnt c [r] []) | traverse cnt ((c as Const (@{const_name Language.com.Seq},_))$c1$c2) = let val (cnt1,v1) = traverse cnt c1; val (cnt2,v2) = traverse cnt1 c2; in (cnt2, f cnt c [] [v1,v2]) end | traverse cnt ((c as Const (@{const_name Language.com.Cond},_))$b$c1$c2) = let val (cnt1,v1) = traverse cnt c1; val (cnt2,v2) = traverse cnt1 c2; in (cnt2, f cnt c [b] [v1,v2]) end | traverse cnt ((c as Const (@{const_name Language.com.While},_))$b$c1) = let val (cnt1,v1) = traverse cnt c1 in (cnt1, f cnt c [b] [v1]) end | traverse cnt ((c as Const (@{const_name Language.com.Call},_))$p) = (cnt, f cnt c [p] []) | traverse cnt ((c as Const (@{const_name Language.com.DynCom},_))$c1) = (cnt, f cnt c [c1] []) | traverse cnt ((c as Const (@{const_name Language.com.Guard},_))$flt$g$c1) = let val (cnt1,v1) = traverse (cnt + length (split_guard g)) c1 in (cnt1, f cnt c [flt,g] [v1]) end | traverse cnt (c as Const (@{const_name Language.com.Throw},_)) = (cnt,f cnt c [] []) | traverse cnt ((c as Const (@{const_name Language.com.Catch},_))$c1$c2) = let val (cnt1,v1) = traverse cnt c1; val (cnt2,v2) = traverse cnt1 c2; in (cnt2, f cnt c [] [v1,v2]) end | traverse cnt ((c as Const (@{const_name Language.guards},_))$gs$c1) = let val (cnt1,v1) = traverse (cnt + length (split_guards gs)) c1; in (cnt1, f cnt c [gs] [v1]) end | traverse cnt ((c as Const (@{const_name Language.block},_))$init$c1$return$c2) = let val (cnt1,v1) = traverse cnt c1 in (cnt1, f cnt c [init,return,c2] [v1]) end | traverse cnt ((c as Const (@{const_name Language.call},_))$init$p$return$c1) = (cnt, f cnt c [init,p,return,c1] []) | traverse cnt ((c as Const (@{const_name Language.whileAnno},_))$b$I$V$c1) = let val (cnt1,v1) = traverse cnt c1 in (cnt1, f cnt c [b,I,V] [v1]) end | traverse cnt ((c as Const (@{const_name Language.whileAnnoG},_))$gs$b$I$V$c1) = let val (cnt1,v1) = traverse (cnt + length (split_guards gs)) c1 in (cnt1, f cnt c [gs,b,I,V] [v1]) end | traverse _ t = raise TERM ("fold_com: unknown command",[t]); in snd (traverse 0 t) end; (*************************** Tactics ****************************************************) (*** Aux. tactics ***) fun cond_rename_bvars cond name thm = let fun rename (tm as (Abs (x, T, t))) = if cond tm then Abs (name x, T, rename t) else Abs (x, T, rename t) | rename (t $ u) = rename t $ rename u | rename a = a; in Thm.renamed_prop (rename (Thm.prop_of thm)) thm end; val rename_bvars = cond_rename_bvars (K true); fun trace_tac ctxt str st = (if Config.get ctxt hoare_trace then tracing str else (); all_tac st); fun error_tac str st = (error str;no_tac st); fun rename_goal ctxt name = EVERY' [K (trace_tac ctxt "rename_goal -- START"), SELECT_GOAL (PRIMITIVE (rename_bvars name)), K (trace_tac ctxt "rename_goal -- STOP")]; (* splits applications of tupled arguments to a schematic Variables, e.g. * ALL a b. ?P (a,b) --> ?Q (a,b) gets * ALL a b. ?P a b --> ?Q a b * only tuples nested to the right are splitted. *) fun split_pair_apps ctxt thm = let val t = Thm.prop_of thm; fun mk_subst subst (Abs (x,T,t)) = mk_subst subst t | mk_subst subst (t as (t1$t2)) = (case strip_comb t of (var as Var (v,vT),args) => (if not (AList.defined (op =) subst var) then let val len = length args; val (argTs,bdyT) = strip_type vT; val (z, _) = Name.variant "z" (fold Term.declare_term_frees args Name.context); val frees = map (apfst (fn i => z^string_of_int i)) (0 upto (len - 1) ~~ argTs); fun splitT (Type (@{type_name Product_Type.prod}, [T1, T2])) = T1::splitT T2 | splitT T = [T]; fun pair_depth (Const (@{const_name Pair},aT)$t1$t2) = 1 + (pair_depth t2) | pair_depth _ = 0; fun mk_sel max free i = let val snds = funpow i HOLogic.mk_snd (Free free) in if i=max then snds else HOLogic.mk_fst snds end; fun split (free,arg) = let val depth = (pair_depth arg); in if depth = 0 then [Free free] else map (mk_sel depth free) (0 upto depth) end; val args' = maps split (frees ~~ args); val argTs' = maps splitT argTs; val inst = fold_rev absfree frees (list_comb (Var (v,argTs' ---> bdyT), args')) in subst@[(var,inst)] end else subst) | _ => mk_subst (mk_subst subst t1) t2) | mk_subst subst t = subst; val subst = map (fn (v,t) => (dest_Var v, Thm.cterm_of ctxt t)) (mk_subst [] t); in full_simplify (put_simpset HOL_basic_ss ctxt addsimps [@{thm fst_conv}, @{thm snd_conv}]) (Drule.instantiate_normalize ([],subst) thm) end; (* Generates split theorems, for !!,!,? quantifiers and for UN, e.g. * ALL x. P x = ALL a b. P a b *) fun mk_split_thms ctxt (vars as _::_) = let val thy = Proof_Context.theory_of ctxt; val names = map fst vars; val types = map snd vars; val free_vars = map Free vars; val pT = foldr1 HOLogic.mk_prodT types; val x = (singleton (Name.variant_list names) "x", pT); val xp = foldr1 HOLogic.mk_prod free_vars; val tfree_names = fold Term.add_tfree_names free_vars []; val zeta = TFree (singleton (Name.variant_list tfree_names) "z", Sign.defaultS thy); val split_meta_prop = let val P = Free (singleton (Name.variant_list names) "P", pT --> Term.propT) in Logic.mk_equals (Logic.all (Free x) (P $ Free x), fold_rev Logic.all free_vars (P $ xp)) end; val P = Free (singleton (Name.variant_list names) "P", pT --> HOLogic.boolT); val split_object_prop = let fun ALL vs t = Library.foldr (fn ((v,T),t) => HOLogic.mk_all (v,T,t)) (vs,t) in (ALL [x] (P $ Free x)) === (ALL vars (P $ xp)) end; val split_ex_prop = let fun EX vs t = Library.foldr (fn ((v,T),t) => HOLogic.mk_exists (v,T,t)) (vs,t) in (EX [x] (P $ Free x)) === (EX vars (P $ xp)) end; val split_UN_prop = let val P = Free (singleton (Name.variant_list names) "P", pT --> HOLogic.mk_setT zeta); fun UN vs t = Library.foldr mk_UN (vs, t) in (UN [x] (P $ Free x)) === (UN vars (P $ xp)) end; fun prove_simp simps prop = let val ([prop'], _) = Variable.importT_terms [prop] ctxt (* FIXME continue context!? *) in Goal.prove_global thy [] [] prop' (fn {context = ctxt, ...} => ALLGOALS (simp_tac (put_simpset HOL_basic_ss ctxt addsimps simps))) end; val split_meta = prove_simp [@{thm split_paired_all}] split_meta_prop; val split_object = prove_simp [@{thm split_paired_All}] split_object_prop; val split_ex = prove_simp [@{thm split_paired_Ex}] split_ex_prop; val split_UN = prove_simp [@{thm Hoare.split_paired_UN}] split_UN_prop; in [split_meta,split_object,split_ex,split_UN] end | mk_split_thms _ _ = raise Match; fun rename_aux_var name rule = let fun is_aux_var (Abs ("Z",TVar(_,_),_)) = true | is_aux_var _ = false; in cond_rename_bvars is_aux_var (K name) rule end; (* adapts single auxiliary variable in a rule to potentialy multiple auxiliary * variables in actual specification, e.g. if vars are a b, * split_app=false: ALL Z. ?P Z gets to ALL a b. ?P (a,b) * split_app=true: ALL Z. ?P Z gets to ALL a b. ?P a b * If only one auxiliary variable is given, the variables are just renamed, * If no auxiliary is given, unit is inserted for Z: * ALL Z. ?P Z gets P () *) fun adapt_aux_var ctxt split_app (vars as (_::_::_)) tvar_rules = let val thy = Proof_Context.theory_of ctxt; val max_idx = fold Integer.max (map (Thm.maxidx_of o snd) tvar_rules) 0; val types = map (fn i => TVar (("z",i),Sign.defaultS thy)) (max_idx + 1 upto (max_idx + length vars)); fun tvar n = (n, Sign.defaultS thy); val pT = Thm.ctyp_of ctxt (foldr1 HOLogic.mk_prodT types); val rules' = map (fn (z,r) => (Drule.instantiate_normalize ([(tvar z,pT)],[]) r)) tvar_rules; val splits = mk_split_thms ctxt (vars ~~ types); val rules'' = map (full_simplify (put_simpset HOL_basic_ss ctxt addsimps splits)) rules'; in if split_app then (map (split_pair_apps ctxt) rules'') else rules'' end | adapt_aux_var _ _ ([name]) tvar_rules = map (rename_aux_var name o snd) tvar_rules | adapt_aux_var ctxt _ ([]) tvar_rules = let val thy = Proof_Context.theory_of ctxt; fun tvar n = (n, Sign.defaultS thy); val uT = Thm.ctyp_of ctxt HOLogic.unitT; val rules' = map (fn (z,r) => (Drule.instantiate_normalize ([(tvar z,uT)],[]) r)) tvar_rules; val splits = [@{thm Hoare.unit_meta},@{thm Hoare.unit_object},@{thm Hoare.unit_ex},@{thm Hoare.unit_UN}]; val rules'' = map (full_simplify (put_simpset HOL_basic_ss ctxt addsimps splits)) rules'; in rules'' end (* Generates a rule for recursion for n procedures out of general recursion rule *) fun gen_call_rec_rule ctxt specs_name n rule = let val thy = Proof_Context.theory_of ctxt; val maxidx = Thm.maxidx_of rule; val vars = Term.add_vars (Thm.prop_of rule) []; fun get_type n = the (AList.lookup (op =) vars (n, 0)); val (Type (_, [Type (_, [assT, Type (_, [pT,_])])])) = get_type specs_name; val zT = TVar (("z",maxidx+1),Sign.defaultS thy) fun mk_var i n T = Var ((n ^ string_of_int i,0),T); val quadT = HOLogic.mk_prodT (assT, HOLogic.mk_prodT (pT, HOLogic.mk_prodT (assT,assT))); val quadT_set = HOLogic.mk_setT quadT; fun mk_spec i = let val quadruple = HOLogic.mk_ptuple (HOLogic.flat_tupleT_paths quadT) quadT [mk_var i "P" (zT --> assT)$Bound 0, mk_var i "p" pT, mk_var i "Q" (zT --> assT)$Bound 0, mk_var i "A" (zT --> assT)$Bound 0]; val single = HOLogic.mk_set quadT [quadruple]; in mk_UN' zT quadT_set (Abs ("Z", zT, single)) end; val Specs = foldr1 (mk_Un quadT_set) (map mk_spec (1 upto n)); val rule' = Thm.instantiate' [] [SOME (Thm.cterm_of ctxt Specs)] rule |> full_simplify (put_simpset (simpset_of @{theory_context Main}) ctxt addsimps [@{thm Hoare.conjE_simp},@{thm Hoare.in_Specs_simp},@{thm Hoare.in_set_Un_simp},@{thm split_all_conj}, @{thm image_Un},@{thm image_Un_single_simp}] ) |> rename_bvars (fn s => if member (op =) ["s","\"] s then s else "Z") in rule' end; fun gen_proc_rec ctxt mode n = gen_call_rec_rule ctxt "Specs" n (ProcRecSpecs mode); (*** verification condition generator ***) (* simplifications on "Collect" sets, like {s. P s} Int {s. Q s} = {s. P s & Q s} *) fun merge_assertion_simp_tac ctxt thms = simp_tac (put_simpset HOL_basic_ss ctxt addsimps ([@{thm Hoare.CollectInt_iff},@{thm HOL.conj_assoc},@{thm Hoare.Compl_Collect},singleton_conv_sym, @{thm Set.Int_empty_right},@{thm Set.Int_empty_left},@{thm Un_empty_right},@{thm Un_empty_left}]@thms)) ; (* The following probably shouldn't live here, but refactoring so that Hoare could depend on recursive_records does not look feasible. The upshot is that there's a duplicate foldcong_ss set here. *) structure FoldCongData = Theory_Data ( type T = simpset; val empty = HOL_basic_ss; val copy = I; val extend = I; val merge = merge_ss; ) val get_foldcong_ss = FoldCongData.get fun add_foldcongs congs thy = FoldCongData.map (fn ss => Proof_Context.init_global thy |> put_simpset ss |> fold Simplifier.add_cong congs |> simpset_of) thy fun add_foldcongsimps simps thy = FoldCongData.map (fn ss => Proof_Context.init_global thy |> put_simpset ss |> (fn ctxt => ctxt addsimps simps) |> simpset_of) thy (* propagates state into "Collect" sets and simplifies selections updates like: * s:{s. P s} = P s *) fun in_assertion_simp_tac ctxt state_kind thms i = let val vcg_simps = #vcg_simps (get_data ctxt); val fold_simps = get_foldcong_ss (Proof_Context.theory_of ctxt) in EVERY [simp_tac (put_simpset HOL_basic_ss ctxt addsimps ([mem_Collect_eq,@{thm Set.Un_iff},@{thm Set.Int_iff}, @{thm Set.empty_subsetI}, @{thm Set.empty_iff}, UNIV_I, @{thm Hoare.Collect_False}]@thms@K_convs@vcg_simps) addsimprocs (state_simprocs state_kind) |> fold Simplifier.add_cong K_congs) i THEN_MAYBE (simp_tac (put_simpset fold_simps ctxt addsimprocs [state_upd_simproc state_kind]) i) ] end; fun assertion_simp_tac ctxt state_kind thms i = merge_assertion_simp_tac ctxt [] i THEN_MAYBE in_assertion_simp_tac ctxt state_kind thms i (* simplify equality test on strings (and datatype-constructors) and propagate result*) fun string_eq_simp_tac ctxt = simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms list.inject list.distinct char.inject cong_exp_iff_simps simp_thms}); fun assertion_string_eq_simp_tac ctxt state_kind thms i = assertion_simp_tac ctxt state_kind thms i THEN_MAYBE string_eq_simp_tac ctxt i; fun before_set2pred_simp_tac ctxt = (simp_tac (put_simpset HOL_basic_ss ctxt addsimps [singleton_conv_sym, @{thm Hoare.CollectInt_iff}, @{thm Hoare.Compl_Collect}])); (*****************************************************************************) (** set2pred transforms sets inclusion into predicates implication, **) (** maintaining the original variable names. **) (** Ex. "{x. x=0} <= {x. x <= 1}" -set2pred-> "x=0 --> x <= 1" **) (** Subgoals containing intersections (A Int B) or complement sets (-A) **) (** are first simplified by "before_set2pred_simp_tac", that returns only **) (** subgoals of the form "{x. P x} <= {x. Q x}", which are easily **) (** transformed. **) (** This transformation may solve very easy subgoals due to a ligth **) (** simplification done by full_simp_tac **) (*****************************************************************************) fun set2pred_tac ctxt i thm = ((before_set2pred_simp_tac ctxt i) THEN_MAYBE (EVERY [trace_tac ctxt "set2pred", resolve_tac ctxt [subsetI] i, resolve_tac ctxt [CollectI] i, dresolve_tac ctxt [CollectD] i, full_simp_tac (put_simpset HOL_basic_ss ctxt) i ])) thm (*****************************************************************************) (** BasicSimpTac is called to simplify all verification conditions. It does **) (** a light simplification by applying "mem_Collect_eq" **) (** then it tries to solve subgoals of the form "A <= A" and then if **) (** set2pred is true it **) (** transforms any other into predicates, applying then **) (** the tactic chosen by the user, which may solve the subgoal completely **) (** (MaxSimpTac). **) (*****************************************************************************) fun MaxSimpTac ctxt tac i = TRY (FIRST[resolve_tac ctxt [subset_refl] i, set2pred_tac ctxt i THEN_MAYBE tac i, trace_tac ctxt "final_tac failed" ]); fun BasicSimpTac ctxt state_kind set2pred thms tac i = EVERY [(trace_tac ctxt "BasicSimpTac -- START --"), assertion_simp_tac ctxt state_kind thms i THEN_MAYBE (if set2pred then MaxSimpTac ctxt tac i else TRY (resolve_tac ctxt [subset_refl] i)), (trace_tac ctxt "BasicSimpTac -- STOP --")]; (* EVERY [(trace_tac ctxt "BasicSimpTac -- START --"), simp_tac (HOL_basic_ss addsimps [mem_Collect_eq,@{thm Hoare.CollectInt_iff}, @{thm Set.empty_subsetI}, @{thm Set.empty_iff}, UNIV_I] addsimprocs [state_simproc sk]) i THEN_MAYBE simp_tac (HOL_basic_ss addsimprocs [state_upd_simproc sk]) i THEN_MAYBE (if set2pred then MaxSimpTac ctxt tac i else TRY (rtac subset_refl i)), (trace_tac ctxt "BasicSimpTac -- STOP --")]; *) (* fun simp_state_eq_tac Record state_space = full_simp_tac (HOL_basic_ss addsimprocs (state_simprocs Record)) THEN_MAYBE' full_simp_tac (HOL_basic_ss addsimprocs [state_upd_simproc Record]) THEN_MAYBE' (state_split_simp_tac [] state_space) | simp_state_eq_tac StateFun state_space = *) fun post_conforms_tac ctxt state_kind i = EVERY [REPEAT1 (resolve_tac ctxt [allI,impI] i), ((fn i => TRY (resolve_tac ctxt [conjI] i)) THEN_ALL_NEW (fn i => (REPEAT (resolve_tac ctxt [allI,impI] i)) THEN (full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [mem_Collect_eq,@{thm Set.singleton_iff},@{thm Set.empty_iff},UNIV_I] addsimprocs (state_simprocs state_kind)) i))) i]; fun dest_hoare_raw (Const(@{const_name HoarePartialDef.hoarep},_)$G$T$F$P$C$Q$A) = (P,C,Q,A,Partial,G,T,F) | dest_hoare_raw (Const(@{const_name HoareTotalDef.hoaret},_)$G$T$F$P$C$Q$A) = (P,C,Q,A,Total,G,T,F) | dest_hoare_raw t = raise TERM ("Hoare.dest_hoare_raw: unexpected term",[t]) fun mk_hoare_abs Ts (P,C,Q,A,mode,G,T,F) = let val hoareT = map (curry fastype_of1 Ts) [G,T,F,P,C,Q,A] ---> HOLogic.boolT; val hoareC = (case mode of Partial => Const (@{const_name HoarePartialDef.hoarep},hoareT) | Total => Const (@{const_name HoareTotalDef.hoaret},hoareT)); in hoareC$G$T$F$P$C$Q$A end; val is_hoare = can dest_hoare_raw fun dest_hoare t = let val triple = (strip_qnt_body @{const_name "All"} o HOLogic.dest_Trueprop o strip_qnt_body @{const_name Pure.all}) t; in dest_hoare_raw triple end; fun get_aux_tvar rule = let fun aux_hoare (Abs ("Z",TVar (z,_),t)) = if is_hoare (strip_qnt_body @{const_name All} t) then SOME z else NONE | aux_hoare _ = NONE; in (case first_subterm_dest (aux_hoare) (Thm.prop_of rule) of SOME (_,z) => (z,rule) | NONE => raise TERM ("get_aux_tvar: No auxiliary variable of hoare-rule found", [Thm.prop_of rule])) end; fun strip_vars t = let val bdy = (HOLogic.dest_Trueprop o Logic.strip_assums_concl) t; in strip_qnt_vars @{const_name Pure.all} t @ strip_qnt_vars @{const_name All} bdy end; local (* ex_simps are necessary in case of multiple logical variables. The state will usually be the first variable. EX s a b. s=s' ... . We have to transport EX s to s=s' to perform the substitution *) val conseq1_ss_base = simpset_of (put_simpset HOL_basic_ss @{context} addsimps ([mem_Collect_eq,@{thm Set.singleton_iff},@{thm Set.Int_iff}, @{thm Set.empty_iff},UNIV_I, @{thm HOL.conj_assoc}, @{thm disj_assoc}] @ @{thms Hoare.all_imp_eq_triv} @K_convs @ @{thms simp_thms} @ @{thms ex_simps} @ @{thms all_simps}) delsimps [@{thm Hoare.all_imp_to_ex}] |> fold Simplifier.add_cong K_congs) val conseq1_ss_record = simpset_of (put_simpset conseq1_ss_base @{context} addsimprocs (state_simprocs Record)); val conseq1_ss_fun = simpset_of (put_simpset conseq1_ss_base @{context} addsimprocs (state_simprocs Function)); fun conseq1_ss Record = conseq1_ss_record | conseq1_ss Function = conseq1_ss_fun; val conseq2_ss_base = simpset_of (put_simpset HOL_basic_ss @{context} addsimps (@{thms Hoare.all_imp_eq_triv} @ @{thms simp_thms} @ @{thms ex_simps} @ @{thms all_simps}) delsimps [@{thm Hoare.all_imp_to_ex}] |> Simplifier.add_cong @{thm imp_cong}); val conseq2_ss_record = simpset_of (put_simpset conseq2_ss_base @{context} addsimprocs [state_upd_simproc Record, state_ex_sel_eq_simproc Record]); val conseq2_ss_fun = simpset_of (put_simpset conseq2_ss_base @{context} addsimprocs [state_upd_simproc Function, state_ex_sel_eq_simproc Function]); fun conseq2_ss Record = conseq2_ss_record | conseq2_ss Function = conseq2_ss_fun; in fun raw_conseq_simp_tac ctxt state_kind thms i = let val ctxt' = Config.put simp_depth_limit 0 ctxt; in simp_tac (put_simpset (conseq1_ss state_kind) ctxt' addsimps thms) i THEN_MAYBE simp_tac (put_simpset (conseq2_ss state_kind) ctxt') i end end val conseq_simp_tac = raw_conseq_simp_tac; (* Generates the hoare-quadruples that can be derived out of the hoare-context T *) fun gen_context_thms ctxt mode params G T F = let val Type (_,[comT]) = range_type (fastype_of G); fun destQuadruple (Const (@{const_name Set.insert},_) $ PpQA $ Const (@{const_name Orderings.bot}, _)) = PpQA | destQuadruple t = raise Match; fun mkCallQuadruple (Const (@{const_name Pair}, _) $ P $ (Const (@{const_name Pair}, _) $ p $ (Const(@{const_name Pair}, _) $ Q $ A))) = let val Call_p = Const (@{const_name Language.com.Call}, fastype_of p --> comT) $ p; in (P, Call_p, Q, A) end; fun mkHoare mode G T F (vars,PpQA) = let val hoare = (case mode of Partial => @{const_name HoarePartialDef.hoarep} | Total => @{const_name HoareTotalDef.hoaret}); (* FIXME: Use future Proof_Context.rename_vars or make closed term and remove by hand *) (* fun free_params ps t = foldr (fn ((x,xT),t) => snd (variant_abs (x,xT,t))) (ps,t); val PpQA' = mkCallQuadruple (strip_qnt_body @{const_name Pure.all} (free_params params (Term.list_all (vars,PpQA)))); *) val params' = (Variable.variant_frees ctxt [PpQA] params); val bnds = map Bound (0 upto (length vars - 1)); fun free_params_vars t = subst_bounds (bnds @ rev (map Free params' ), t) fun free_params t = subst_bounds (rev (map Free params' ), t) val (P',p',Q',A') = mkCallQuadruple (free_params_vars PpQA); val T' = free_params T; val G' = free_params G; val F' = free_params F; val bdy = mk_hoare_abs [] (P',p',Q',A',mode,G',T',F'); in (HOLogic.mk_Trueprop (HOLogic.list_all (vars,bdy)), map fst params') end; fun hoare_context_specs mode G T F = let fun mk t = try (mkHoare mode G T F o apsnd destQuadruple o dest_UN) t; in map_filter mk (dest_Un T) end; fun mk_prove mode (prop,params) = let val vars = map fst (strip_qnt_vars @{const_name All} (HOLogic.dest_Trueprop (Logic.strip_assums_concl prop))); val [asmUN'] = adapt_aux_var ctxt true vars [get_aux_tvar (AsmUN mode)] in Goal.prove ctxt params [] prop (fn thms => EVERY[trace_tac ctxt "extracting specifications from hoare context", resolve_tac ctxt [asmUN'] 1, DEPTH_SOLVE_1 (resolve_tac ctxt [subset_refl,refl] 1 ORELSE ((resolve_tac ctxt [@{thm Hoare.subset_unI1}] 1 APPEND resolve_tac ctxt [@{thm Hoare.subset_unI2}] 1) ORELSE (resolve_tac ctxt [@{thm Hoare.subset_singleton_insert1}] 1 APPEND resolve_tac ctxt [@{thm Hoare.subset_singleton_insert2}] 1))) ORELSE error_tac ("vcg: cannot extract specifications from context") ]) end; val specs = hoare_context_specs mode G T F; in map (mk_prove mode) specs end; fun is_modifies_clause t = exists_subterm (fn (Const (@{const_name Hoare.meq},_)) => true| _ => false) (#3 (dest_hoare (Logic.strip_assums_concl t))) handle (TERM _) => false; val is_spec_clause = not o is_modifies_clause; (* e.g: Intg => the_Intg lift Intg => lift the_Intg map Ingt => map the_Intg Hp o lift Intg => lift the_Intg o the_Hp *) fun swap_constr_destr f (t as (Const (@{const_name Fun.id},_))) = t | swap_constr_destr f (t as (Const (c,Type ("fun",[T,valT])))) = (Const (f c, Type ("fun",[valT,T])) handle Empty => raise TERM ("Hoare.swap_constr_destr",[t])) | swap_constr_destr f (Const ("StateFun.map_fun",Type ("fun", (* FIXME unknown "StateFun.map_fun" !? *) [Type ("fun",[T,valT]), Type ("fun",[Type ("fun",[xT,T']), Type ("fun",[xT',valT'])])]))$g) = Const ("StateFun.map_fun",Type("fun",[Type ("fun",[valT,T]), Type ("fun",[Type ("fun",[xT,valT']), Type ("fun",[xT',T'])])]))$ swap_constr_destr f g | swap_constr_destr f (Const (@{const_name Fun.comp},Type ("fun", [Type ("fun",[bT',cT]), Type ("fun",[Type ("fun",[aT ,bT]), Type ("fun",[aT',cT'])])]))$h$g) = let val h'=swap_constr_destr f h; val g'=swap_constr_destr f g; in Const (@{const_name Fun.comp},Type ("fun", [Type ("fun",[bT,aT]), Type ("fun",[Type ("fun",[cT,bT']), Type ("fun",[cT',aT'])])]))$g'$h' end | swap_constr_destr f (Const (@{const_name List.map},Type ("fun", [Type ("fun",[aT,bT]), Type ("fun",[asT,bsT])]))$g) = (Const (@{const_name List.map},Type ("fun", [Type ("fun",[bT,aT]), Type ("fun",[bsT,asT])]))$swap_constr_destr f g) | swap_constr_destr f t = raise TERM ("Hoare.swap_constr_destr",[t]); (* FIXME: unused? *) val destr_to_constr = let fun convert c = let val (path,base) = split_last (Long_Name.explode c); in Long_Name.implode (path @ ["val",unprefix "the_" base]) end; in swap_constr_destr convert end; fun gen_call_tac cont_tac mode cmode state_kind state_space ctxt asms spec_sfx pname return has_args _ = let val thy = Proof_Context.theory_of ctxt; val pname' = unsuffix proc_deco pname; val spec = (case AList.lookup (op =) asms pname of SOME s => SOME s | NONE => try (Proof_Context.get_thm ctxt) (suffix spec_sfx pname')); fun auxvars_for p t = (case first_subterm_dest (try dest_call) t of SOME (vars,(_,p',_,_,m,_)) => (if m=Static andalso p=(dest_string' p') then SOME vars else NONE) | NONE => NONE); fun get_auxvars_for p t = (case (map_filter ((auxvars_for p) o snd) (max_subterms_dest tap_UN t)) of (vars::_) => vars | _ => []); fun spec_tac augment_rule augment_emptyFaults _ spec i = let val spec' = augment_emptyFaults OF [spec] handle THM _ => spec; in EVERY [resolve_tac ctxt [augment_rule] i, resolve_tac ctxt [spec'] (i+1), TRY (resolve_tac ctxt [subset_refl, @{thm Set.empty_subsetI}, @{thm Set.Un_upper1}, @{thm Set.Un_upper2}] i)] end; fun check_spec name P thm = (case try dest_hoare (Thm.concl_of thm) of SOME spc => (case try dest_call (#2 spc) of SOME (_,p,_,_,m,_) => if proc_name m p = name andalso P (Thm.concl_of thm) then SOME (#5 spc,thm) else NONE | _ => NONE) | _ => NONE) fun find_dyn_specs name P thms = map_filter (check_spec name P) thms; fun get_spec name P thms = case find_dyn_specs name P thms of (spec_mode,spec)::_ => SOME (spec_mode,spec) | _ => NONE; fun solve_spec augment_rule _ augment_emptyFaults mode _ (SOME spec_mode) (SOME spec) i= if mode=Partial andalso spec_mode=Total then resolve_tac ctxt [@{thm HoareTotal.hoaret_to_hoarep'}] i THEN spec_tac augment_rule augment_emptyFaults mode spec i else if mode=spec_mode then spec_tac augment_rule augment_emptyFaults mode spec i else error("vcg: cannot use a partial correctness specification of " ^ pname' ^ " for a total correctness proof") | solve_spec _ asmUN_rule _ mode Static _ _ i =(* try to infer spec out of context *) EVERY[trace_tac ctxt "inferring specification from hoare context1", resolve_tac ctxt [asmUN_rule] i, DEPTH_SOLVE_1 (resolve_tac ctxt [subset_refl,refl] i ORELSE ((resolve_tac ctxt [@{thm Hoare.subset_unI1}] i APPEND resolve_tac ctxt [@{thm Hoare.subset_unI2}] i) ORELSE (resolve_tac ctxt [@{thm Hoare.subset_singleton_insert1}] i APPEND resolve_tac ctxt [@{thm Hoare.subset_singleton_insert2}] i))) ORELSE error_tac ("vcg: cannot infer specification of " ^ pname' ^ " from context") (* if tactic for DEPTH_SOLVE_1 would create new subgoals, use SELECT_GOAL and DEPTH_SOLVE *) ] | solve_spec augment_rule asmUN_rule augment_emptyFaults mode Parameter _ _ i = (* try to infer spec out of assumptions *) let fun tac thms = (case (find_dyn_specs pname is_spec_clause thms) of (spec_mode,spec)::_ => solve_spec augment_rule asmUN_rule augment_emptyFaults mode Parameter (SOME spec_mode) (SOME spec) 1 | _ => all_tac) in Subgoal.FOCUS (tac o #prems) ctxt i end val strip_spec_vars = strip_qnt_vars @{const_name All} o HOLogic.dest_Trueprop; fun apply_call_tac ctxt pname mode cmode spec_mode spec_goal is_abr spec (subgoal,i) = let val spec_vars = map fst (case spec of SOME sp => (strip_spec_vars (Thm.concl_of sp)) | NONE => (case try (dest_hoare) subgoal of SOME (_,_,_,_,_,_,Theta,_) => get_auxvars_for pname Theta | _ => [])); fun get_call_rule Static mode is_abr = if is_abr then Proc mode else ProcNoAbr mode | get_call_rule Parameter mode is_abr = if is_abr then DynProcProcPar mode else DynProcProcParNoAbr mode; val [call_rule,augment_ctxt_rule,asmUN_rule, augment_emptyFaults] = adapt_aux_var ctxt true spec_vars (map get_aux_tvar [get_call_rule cmode mode is_abr, AugmentContext mode, AsmUN mode, AugmentEmptyFaults mode]); in EVERY [resolve_tac ctxt [call_rule] i, trace_tac ctxt "call_tac -- basic_tac -- solving spec", solve_spec augment_ctxt_rule asmUN_rule augment_emptyFaults mode cmode spec_mode spec spec_goal] end; fun basic_tac spec i = let val msg ="Theorem " ^pname'^spec_sfx ^ " is no proper specification for procedure " ^pname'^ "; trying to infer specification from hoare context"; fun spec' s mode abr = if is_modifies_clause (Thm.concl_of s) then if abr then (TrivPost mode) OF [s] else (TrivPostNoAbr mode) OF [s] else s; val (is_abr,spec_mode,spec,spec_has_args) = (* is there a proper specification fact? *) case spec of NONE => (true,NONE,NONE,false) | SOME s => case try dest_hoare (Thm.concl_of s) of NONE => (warning msg;(true,NONE,NONE,false)) | SOME (_,c,Q,spec_abr,spec_mode,_,_,_) => case try dest_call c of NONE => (warning msg;(true,NONE,NONE,false)) | SOME (_,p,_,_,m,spec_has_args) => if proc_name m p = pname then if (mode=Total andalso spec_mode=Partial) then (warning msg;(true,NONE,NONE,false)) else if is_empty_set spec_abr then (false,SOME spec_mode, SOME (spec' s spec_mode false),spec_has_args) else (true,SOME spec_mode, SOME (spec' s spec_mode true),spec_has_args) else (warning msg;(true,NONE,NONE,false)); val () = if spec_has_args then error "procedure call in specification must be parameterless!" else (); val spec_goal = i+2; in EVERY[trace_tac ctxt "call_tac -- basic_tac -- START --", SUBGOAL (apply_call_tac ctxt pname mode cmode spec_mode spec_goal is_abr spec) i, resolve_tac ctxt [allI] (i+1), resolve_tac ctxt [allI] (i+1), cont_tac (i+1), trace_tac ctxt "call_tac -- basic_tac -- simplify", conseq_simp_tac ctxt state_kind [@{thm StateSpace.upd_globals_def}] i, trace_tac ctxt "call_tac -- basic_tac -- STOP --"] end; fun get_modifies (Const (@{const_name Collect},_) $ Abs (_,_,m)) = m | get_modifies t = raise TERM ("gen_call_tac.get_modifies",[t]); fun get_updates (Abs (_,_,t)) = get_updates t | get_updates (Const (@{const_name Hoare.mex},_) $ t) = get_updates t | get_updates (Const (@{const_name Hoare.meq},T) $ _ $ upd) = (T,upd) | get_updates t = raise TERM ("gen_call_tac.get_updates",[t]); (* return has the form: %s t. s(|globals:=globals t,...|) * should be translated to %s t. s(|globals := globals s(|m := m (globals t),...|),...|) * for all m in the modifies list. *) fun mk_subst gT meqT = fst (Sign.typ_unify thy (gT,domain_type meqT) (Vartab.empty,0)); fun mk_selR subst gT (upd,uT) = let val vT = range_type (hd (binder_types uT)); in Const (unsuffix Record.updateN upd,gT --> (Envir.norm_type subst vT)) end; (* lookup:: "('v => 'a) => 'n => ('n => 'v) => 'a" * update:: "('v => 'a) => ('a => 'v) => 'n => ('a => 'a) => ('n => 'v) => ('n => 'v)" *) fun mk_selF subst uT d n = let val vT_a::a_vT::nT::aT_aT::n_vT::_ = binder_types uT; val lT = (Envir.norm_type subst (vT_a --> nT --> n_vT --> (domain_type aT_aT))); val d' = map_types (Envir.norm_type subst) d; in Const (@{const_name StateFun.lookup},lT)$d'$n end; fun mk_rupdR subst gT (upd,uT) = let val [vT,_] = binder_types uT in Const (upd,(Envir.norm_type subst vT) --> gT --> gT) end; fun K_fun kn uT = let val T=range_type (hd (binder_types uT)) in Const (kn,T --> T --> T) end; fun K_rec uT t = let val T=range_type (hd (binder_types uT)) in Abs ("_", T, incr_boundvars 1 t) end; fun mk_supdF subst uT d c n = let val uT' = Envir.norm_type subst uT; val c' = map_types (Envir.norm_type subst) c; val d' = map_types (Envir.norm_type subst) d; in Const (@{const_name StateFun.update},uT')$d'$c'$n end; fun modify_updatesR subst gT glob ((Const (upd,uT))$_$(Const _$Z)) = mk_rupdR subst gT (upd,uT)$ (K_rec uT (mk_selR subst gT (upd,uT)$(glob$Bound 0)))$(glob$Bound 1) | modify_updatesR subst gT glob ((Const (upd,uT))$_$s) = mk_rupdR subst gT (upd,uT)$ (K_rec uT (mk_selR subst gT (upd,uT)$(glob$Bound 0)))$ modify_updatesR subst gT glob s | modify_updatesR subst gT glob ((_$Z)) = (glob$Bound 1) (* may_not_modify *) | modify_updatesR _ _ _ t = raise TERM ("gen_call_tac.modify_updatesR",[t]); fun modify_updatesF subst _ glob (Const (@{const_name StateFun.update},uT)$d$c$n$_$(Const globs$Z)) = mk_supdF subst uT d c n$ (K_fun KNF uT$(mk_selF subst uT d n$(glob$Bound 0)))$(glob$Bound 1) | modify_updatesF subst gT glob (Const (@{const_name StateFun.update},uT)$d$c$n$_$s) = mk_supdF subst uT d c n$ (K_fun KNF uT$(mk_selF subst uT d n$(glob$Bound 0)))$modify_updatesF subst gT glob s | modify_updatesF subst _ glob ((globs$Z)) = (glob$Bound 1) (* may_not_modify *) | modify_updatesF _ _ _ t = raise TERM ("gen_call_tac.modify_updatesF",[t]); fun modify_updates Record = modify_updatesR | modify_updates _ = modify_updatesF fun globalsT (Const (gupd,T)) = domain_type (hd (binder_types T)) | globalsT t = raise TERM ("gen_call_tac.globalsT",[t]); fun mk_upd meqT mods (gupd$(Abs (dmy,dmyT,(glob$Bound 1)))$Bound 1) = let val gT = (globalsT gupd); val subst = mk_subst gT meqT; in (gupd$(Abs (dmy,dmyT,incr_boundvars 1 (modify_updates state_kind subst gT glob mods)))$Bound 1) end | mk_upd meqT mods (upd$v$s) = upd$v$mk_upd meqT mods s | mk_upd _ _ t = raise TERM ("gen_call_tac.mk_upd",[t]); fun modify_return (meqT,mods) (Abs (s,T,Abs (t,U,upd))) = (Abs (s,T,Abs (t,U,mk_upd meqT mods upd))) | modify_return _ t = raise TERM ("get_call_tac.modify_return",[t]); fun modify_tac spec modifies_thm i = try (fn () => let val (_,call,modif_spec_nrm,modif_spec_abr,modif_spec_mode,G,Theta,_) = dest_hoare (Thm.concl_of modifies_thm); val is_abr = not (is_empty_set modif_spec_abr); val emptyTheta = is_empty_set Theta; (*val emptyFaults = is_empty_set F;*) val spec_has_args = #6 (dest_call call); val () = if spec_has_args then error "procedure call in modifies-specification must be parameterless!" else (); val (mxprem,ModRet) = (case cmode of Static => (8,if is_abr then if emptyTheta then (ProcModifyReturn mode) else (ProcModifyReturnSameFaults mode) else if emptyTheta then (ProcModifyReturnNoAbr mode) else (ProcModifyReturnNoAbrSameFaults mode)) | Parameter => (9,if is_abr then if emptyTheta then (ProcProcParModifyReturn mode) else (ProcProcParModifyReturnSameFaults mode) else if emptyTheta then (ProcProcParModifyReturnNoAbr mode) else (ProcProcParModifyReturnNoAbrSameFaults mode))); val to_prove_prem = (case cmode of Static => 0 | Parameter => 1); val spec_goal = if is_abr then i + mxprem - 5 else i + mxprem - 6 val mods_nrm = modif_spec_nrm |> get_modifies |> get_updates; val return' = modify_return mods_nrm return; (* val return' = if is_abr then let val mods_abr = modif_spec_abr |> get_modifies |> get_updates; in modify_return mods_abr return end else return;*) val cret = Thm.cterm_of ctxt return'; val (_,_,return'_var,_,_,_) = nth (Thm.prems_of ModRet) to_prove_prem |> dest_hoare |> #2 |> dest_call; val ModRet' = infer_instantiate ctxt [(#1 (dest_Var return'_var), cret)] ModRet; val modifies_thm_partial = if modif_spec_mode = Total then @{thm HoareTotal.hoaret_to_hoarep'} OF [modifies_thm] else modifies_thm; fun solve_modifies_tac i = (clarsimp_tac ((ctxt |> put_claset (claset_of @{theory_context Set}) |> put_simpset (simpset_of @{theory_context Set})) addsimps ([@{thm Hoare.mex_def},@{thm Hoare.meq_def},@{thm StateSpace.upd_globals_def}]@K_convs) addsimprocs (state_upd_simproc Record::(state_simprocs state_kind)) |> fold Simplifier.add_cong K_congs) i) THEN_MAYBE EVERY [trace_tac ctxt "modify_tac: splitting record", state_split_simp_tac ctxt [] state_space i]; val cnt = i + mxprem; in EVERY[trace_tac ctxt "call_tac -- modifies_tac --", resolve_tac ctxt [ModRet'] i, solve_spec (AugmentContext Partial) (AsmUN Partial) (AugmentEmptyFaults Partial) Partial Static (SOME Partial) (SOME modifies_thm_partial) spec_goal, if is_abr then EVERY [trace_tac ctxt "call_tac -- Solving abrupt modifies --", solve_modifies_tac (cnt - 6)] else all_tac, trace_tac ctxt "call_tac -- Solving Modifies --", solve_modifies_tac (cnt - 7), basic_tac spec (cnt - 8), if cmode = Parameter then EVERY [resolve_tac ctxt [subset_refl] (cnt - 8), simp_tac (put_simpset HOL_basic_ss ctxt addsimps (@{thm Hoare.CollectInt_iff} :: @{thms simp_thms})) 1] else all_tac] end) () |> (fn SOME res => res | NONE => raise TERM ("get_call_tac.modify_tac: no proper modifies spec", [])); fun specs_of_assms_tac thms = (case get_spec pname is_spec_clause thms of SOME (_,spec) => (case get_spec pname is_modifies_clause thms of SOME (_,modifies_thm) => modify_tac (SOME spec) modifies_thm 1 | NONE => basic_tac (SOME spec) 1) | NONE => (warning ("no proper specification for procedure " ^pname^ " in assumptions"); all_tac)); val test_modify_in_ctxt_tac = let val mname = (suffix modifysfx pname'); val mspec = (case try (Proof_Context.get_thm ctxt) mname of SOME s => SOME s | NONE => (case AList.lookup (op =) asms pname of SOME s => if is_modifies_clause (Thm.concl_of s) then SOME s else NONE | NONE => NONE)); in (case mspec of NONE => basic_tac spec | SOME modifies_thm => (case check_spec pname is_modifies_clause modifies_thm of SOME _ => modify_tac spec modifies_thm | NONE => (warning ("ignoring theorem " ^ (suffix modifysfx pname') ^ "; no proper modifies specification for procedure "^pname'); basic_tac spec))) end; fun inline_bdy_tac has_args i = (case try (Proof_Context.get_thm ctxt) (suffix bodyP pname') of NONE => no_tac | SOME impl => (case try (Proof_Context.get_thm ctxt) (suffix (body_def_sfx^"_def") pname') of NONE => no_tac | SOME bdy => (tracing ("No specification found for procedure \"" ^ pname' ^ "\". Inlining procedure!"); if has_args then EVERY [trace_tac ctxt "inline_bdy_tac args", resolve_tac ctxt [CallBody mode] i, resolve_tac ctxt [impl] (i+3), resolve_tac ctxt [allI] (i+2), resolve_tac ctxt [allI] (i+2), in_assertion_simp_tac ctxt state_kind [] (i+2), cont_tac (i+2), resolve_tac ctxt [allI] (i+1),in_assertion_simp_tac ctxt state_kind [bdy] (i+1), cont_tac (i+1), in_assertion_simp_tac ctxt state_kind [@{thm StateSpace.upd_globals_def}] i] else EVERY [trace_tac ctxt "inline_bdy_tac no args", resolve_tac ctxt [ProcBody mode] i, resolve_tac ctxt [impl] (i+2), simp_tac (put_simpset HOL_basic_ss ctxt addsimps [bdy]) (i+1), cont_tac (i+1)]))); in (case cmode of Static => if get_recursive pname ctxt = SOME false andalso is_none spec then inline_bdy_tac has_args else test_modify_in_ctxt_tac | Parameter => (case spec of NONE => (tracing "no spec found!"; Subgoal.FOCUS (specs_of_assms_tac o #prems) ctxt) | SOME spec => (tracing "found spec!"; case check_spec pname is_spec_clause spec of SOME _ => test_modify_in_ctxt_tac | NONE => (warning ("ignoring theorem " ^ (suffix spec_sfx pname') ^ "; no proper specification for procedure " ^pname'); Subgoal.FOCUS (specs_of_assms_tac o #prems) ctxt)))) end; fun call_tac cont_tac mode state_kind state_space ctxt asms spec_sfx t = let val (_,c,_,_,_,_,_,F) = dest_hoare (Logic.strip_assums_concl t); fun gen_tac (_,pname,return,c,cmode,has_args) = gen_call_tac cont_tac mode cmode state_kind state_space ctxt asms spec_sfx (proc_name cmode pname) return has_args F; in gen_tac (dest_call c) end handle TERM _ => K no_tac; fun solve_in_Faults_tac ctxt i = resolve_tac ctxt [UNIV_I, @{thm in_insert_hd}] i ORELSE SELECT_GOAL (SOLVE (simp_tac (put_simpset (simpset_of @{theory_context Set}) ctxt) 1)) i; local fun triv_simp ctxt = merge_assertion_simp_tac ctxt [mem_Collect_eq] THEN' simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms} |> fold Simplifier.add_cong [@{thm conj_cong}, @{thm imp_cong}]); (* a guarded while produces stupid things, since the guards are put at the end of the body and in the invariant (rule WhileAnnoG): - guard: g /\ g - guarantee: g --> g *) in fun guard_tac ctxt strip cont_tac mode (t,i) = let val (_,c,_,_,_,_,_,F) = dest_hoare (Logic.strip_assums_concl t); val (_,_,_,doStrip) = dest_Guard c; val guarantees = if strip orelse doStrip then [GuardStrip mode, GuaranteeStrip mode] else [Guarantee mode] fun basic_tac i = EVERY [resolve_tac ctxt [Guard mode, GuaranteeAsGuard mode] i, trace_tac ctxt "Guard", cont_tac (i+1), triv_simp ctxt i] fun guarantee_tac i = EVERY [resolve_tac ctxt guarantees i, solve_in_Faults_tac ctxt (i+2), cont_tac (i+1), triv_simp ctxt i] in if is_empty_set F then EVERY [trace_tac ctxt "Guard: basic_tac", basic_tac i] else EVERY [trace_tac ctxt "Guard: trying guarantee_tac", guarantee_tac i ORELSE basic_tac i] end handle TERM _ => no_tac end; fun wf_tac ctxt = simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm Wellfounded.wf_measure},@{thm Wellfounded.wf_lex_prod},@{thm Wfrec.wf_same_fst}, @{thm Hoare.wf_measure_lex_prod},@{thm Wellfounded.wf_inv_image}]); fun in_rel_simp ctxt = simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm Hoare.in_measure_iff},@{thm Hoare.in_lex_iff},@{thm Hoare.in_mlex_iff},@{thm Hoare.in_inv_image_iff}, @{thm split_conv}]); fun while_annotate_tac ctxt inv i st = let val annotateWhile = Thm.lift_rule (Thm.cprem_of st i) @{thm HoarePartial.reannotateWhileNoGuard}; val lifted_inv = fold_rev Term.abs (Logic.strip_params (Logic.get_goal (Thm.prop_of st) i)) inv; val invVar = (#1 o strip_comb o #3 o dest_whileAnno o #2 o dest_hoare) (List.last (Thm.prems_of annotateWhile)); val annotate = infer_instantiate ctxt [(#1 (dest_Var invVar), Thm.cterm_of ctxt lifted_inv)] annotateWhile; in ((trace_tac ctxt ("annotating While with: " ^ Syntax.string_of_term ctxt lifted_inv )) THEN compose_tac ctxt (false,annotate,1) i) st end; fun cond_annotate_tac ctxt inv mode (_,i) st = let val annotateCond = Thm.lift_rule (Thm.cprem_of st i) (CondInv' mode); val lifted_inv = fold_rev Term.abs (Logic.strip_params (Logic.get_goal (Thm.prop_of st) i)) inv; val invVar = List.last (Thm.prems_of annotateCond) |> dest_hoare |> #3 |> strip_comb |> #1; val annotate = infer_instantiate ctxt [(#1 (dest_Var invVar), Thm.cterm_of ctxt lifted_inv)] annotateCond; in ((trace_tac ctxt ("annotating Cond with: "^ Syntax.string_of_term ctxt lifted_inv)) THEN compose_tac ctxt (false,annotate,5) i) st end; fun basic_while_tac ctxt state_kind cont_tac tac mode i = let fun common_tac i = EVERY[if mode=Total then wf_tac ctxt (i+3) else all_tac, BasicSimpTac ctxt state_kind true [] tac (i+2), if mode=Total then in_rel_simp ctxt (i+1) THEN (resolve_tac ctxt [allI] (i+1)) else all_tac, cont_tac (i+1) ] fun basic_tac i = EVERY [resolve_tac ctxt [While mode] i, common_tac i] in EVERY [trace_tac ctxt "basic_while_tac: basic_tac", basic_tac i] end; fun while_tac ctxt state_kind inv cont_tac tac mode t i= let val basic_tac = basic_while_tac ctxt state_kind cont_tac tac mode; in (case inv of NONE => basic_tac i | SOME I => EVERY [while_annotate_tac ctxt I i, basic_tac i]) end handle TERM _ => no_tac fun dest_split (Abs (x,T,t)) = let val (vs,recomb,bdy) = dest_split t; in ((x,T)::vs,fn t' => Abs (x,T,recomb t'),bdy) end | dest_split (c as Const (@{const_name case_prod},_)$Abs(x,T,t)) = let val (vs,recomb,bdy) = dest_split t; in ((x,T)::vs,fn t' => c$Abs (x,T,recomb t'),bdy) end | dest_split t = ([],I,t); fun whileAnnoG_tac ctxt strip_guards mode t i st = let val (_,c,_,_,_,_,_,F) = dest_hoare (Logic.strip_assums_concl t); val (SOME grds,_,I,_,_,fix) = dest_whileAnno c; val Rule = if fix then WhileAnnoGFix mode else WhileAnnoG mode; fun extract_faults (Const (@{const_name Set.insert}, _) $ t $ _) = [t] | extract_faults _ = []; fun leave_grd fs (Const (@{const_name Pair}, _) $ f $ g) = if member (op =) fs f andalso strip_guards then NONE else SOME g | leave_grd fs (Const (@{const_name Language.guaranteeStripPair}, _) $ f $ g) = if member (op =) fs f then NONE else SOME g | leave_grd fs _ = NONE; val (I_vs,I_recomb,I') = dest_split I; val grds' = map_filter (leave_grd (extract_faults F)) (HOLogic.dest_list grds); val pars = (Logic.strip_params (Logic.get_goal (Thm.prop_of st) i)); val J = fold_rev Term.abs pars (I_recomb (fold_rev (mk_Int (map snd (pars@I_vs))) grds' I')); val WhileG = Thm.lift_rule (Thm.cprem_of st i) Rule; val invVar = (fst o strip_comb o #3 o dest_whileAnno o (fn xs => nth xs 1) o snd o strip_comb o #2 o dest_hoare) (List.last (Thm.prems_of WhileG)); val WhileGInst = infer_instantiate ctxt [(#1 (dest_Var invVar), Thm.cterm_of ctxt J)] WhileG; in ((trace_tac ctxt ("WhileAnnoG, adding guards to invariant: " ^ Syntax.string_of_term ctxt J)) THEN compose_tac ctxt (false,WhileGInst,1) i) st end handle TERM _ => no_tac st | Bind => no_tac st (* renames bound state variable according to name given in goal, * before rule specAnno is applied, and solves sidecondition *) fun gen_Anno_tac dest rules tac cont_tac ctxt state_kind (t,i) st = let val vars = (dest o #2 o dest_hoare) (Logic.strip_assums_concl t); val rules' = (case (List.filter (not o null) (map dest_splits vars)) of [] => rules |(xs::_) => adapt_aux_var ctxt false (map fst xs) (map get_aux_tvar rules)); in EVERY [resolve_tac ctxt rules' i, tac, simp_tac (put_simpset HOL_basic_ss ctxt addsimps ([@{thm split_conv}, refl, @{thm Hoare.triv_All_eq}]@anno_defs) addsimprocs [@{simproc case_prod_beta}]) (i+2), simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm split_conv}]) (i+1), simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm split_conv}]) i, REPEAT (resolve_tac ctxt [allI] (i+1)), cont_tac (i+1), conseq_simp_tac ctxt state_kind [] i] st end handle TERM _ => no_tac st; fun specAnno_tac ctxt state_kind cont_tac mode = let fun dest_specAnno (Const (@{const_name Language.specAnno},_)$P$c$Q$A) = [P,c,Q,A] | dest_specAnno t = raise TERM ("dest_specAnno",[t]); val rules = [SpecAnnoNoAbrupt mode,SpecAnno mode]; in gen_Anno_tac dest_specAnno rules all_tac cont_tac ctxt state_kind end; fun whileAnnoFix_tac ctxt state_kind cont_tac mode (t,i) = let fun dest (Const (@{const_name Language.whileAnnoFix},_)$b$I$V$c) = [I,V,c] | dest t = raise TERM ("dest_whileAnnoFix",[t]); val rules = [WhileAnnoFix mode]; fun wf_tac' i = EVERY [REPEAT (resolve_tac ctxt [allI] i), simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm split_conv}]) i, wf_tac ctxt i]; val tac = if mode=Total then EVERY [wf_tac' (i+3),in_rel_simp ctxt (i+1)] else all_tac in gen_Anno_tac dest rules tac cont_tac ctxt state_kind (t,i) end; fun lemAnno_tac ctxt state_kind mode (t,i) st = let fun dest_name (Const (x,_)) = x | dest_name (Free (x,_)) = x | dest_name t = raise TERM ("dest_name",[t]); fun dest_lemAnno ctxt (Const (@{const_name Language.lem},_)$n$c) = let val x = Long_Name.base_name (dest_name n); in (case try (Proof_Context.get_thm ctxt) x of NONE => error ("No lemma: '" ^ x ^ "' found.") | SOME spec => (strip_qnt_vars @{const_name All} (HOLogic.dest_Trueprop (Thm.concl_of spec)),spec)) end | dest_lemAnno _ t = raise TERM ("dest_lemAnno",[t]); val (vars,spec) = dest_lemAnno ctxt (#2 (dest_hoare t)); val rules = [LemAnnoNoAbrupt mode,LemAnno mode]; val rules' = (case vars of [] => rules | xs => adapt_aux_var ctxt true (map fst xs) (map get_aux_tvar rules)); in EVERY [resolve_tac ctxt rules' i, resolve_tac ctxt [spec] (i+1), conseq_simp_tac ctxt state_kind [] i] st end handle TERM _ => no_tac st; fun prems_tac ctxt i = TRY (resolve_tac ctxt (Assumption.all_prems_of ctxt) i); fun mk_proc_assoc thms = let fun name (_,p,_,_,cmode,_) = proc_name cmode p; fun proc_name thm = thm |> Thm.concl_of |> dest_hoare |> #2 |> dest_call |> name; in map (fn thm => (proc_name thm,thm)) thms end; fun mk_hoare_tac cont ctxt mode i (name,tac) = EVERY [trace_tac ctxt ("trying: " ^ name),tac cont ctxt mode i]; (* the main hoare tactic *) fun HoareTac annotate_inv exspecs strip_guards mode state_kind state_space spec_sfx ctxt tac st = let val (P,c,Q,A,_,G,T,F) = dest_hoare (Logic.strip_assums_concl (Logic.get_goal (Thm.prop_of st) 1)); val wp_tacs = #wp_tacs (get_data ctxt); val hoare_tacs = #hoare_tacs (get_data ctxt); val params = (strip_vars (Logic.get_goal (Thm.prop_of st) 1)); val Inv = (if annotate_inv then (* Take the postcondition of the triple as invariant for all *) (* while loops (makes sense for the modifies clause) *) SOME Q else NONE); val exspecthms = map (Proof_Context.get_thm ctxt) exspecs; val asms = try (fn () => mk_proc_assoc (gen_context_thms ctxt mode params G T F @ exspecthms)) () |> the_default []; fun while_annoG_tac (t,i) = whileAnnoG_tac ctxt (annotate_inv orelse strip_guards) mode t i; fun WlpTac tac i = (* WlpTac does not end with subset_refl *) FIRST ([EVERY [resolve_tac ctxt [Seq mode] i,trace_tac ctxt "Seq",HoareRuleTac tac false (i+1)], EVERY [resolve_tac ctxt [Catch mode] i,trace_tac ctxt "Catch",HoareRuleTac tac false (i+1)], EVERY [resolve_tac ctxt [CondCatch mode] i,trace_tac ctxt "CondCatch",HoareRuleTac tac false (i+1)], EVERY [resolve_tac ctxt [BSeq mode] i,trace_tac ctxt "BSeq",HoareRuleTac tac false (i+1)], EVERY [resolve_tac ctxt [FCall mode] i,trace_tac ctxt "FCall"], EVERY [resolve_tac ctxt [GuardsNil mode] i,trace_tac ctxt "GuardsNil"], EVERY [resolve_tac ctxt [GuardsConsGuaranteeStrip mode] i, trace_tac ctxt "GuardsConsGuaranteeStrip"], EVERY [resolve_tac ctxt [GuardsCons mode] i,trace_tac ctxt "GuardsCons"], EVERY [SUBGOAL while_annoG_tac i] ] @ map (mk_hoare_tac (HoareRuleTac tac) ctxt mode i) wp_tacs) and HoareRuleTac tac pre_cond i st = let fun call (t,i) = call_tac (HoareRuleTac tac false) mode state_kind state_space ctxt asms spec_sfx t i fun cond_tac i = if annotate_inv andalso Config.get ctxt use_cond_inv_modifies then EVERY[SUBGOAL (cond_annotate_tac ctxt (the Inv) mode) i, HoareRuleTac tac false (i+4), HoareRuleTac tac false (i+3), BasicSimpTac ctxt state_kind true [] tac (i+2), BasicSimpTac ctxt state_kind true [] tac (i+1) ] else EVERY[resolve_tac ctxt [Cond mode] i,trace_tac ctxt "Cond", HoareRuleTac tac false (i+2), HoareRuleTac tac false (i+1)]; fun switch_tac i = EVERY[resolve_tac ctxt [SwitchNil mode] i, trace_tac ctxt "SwitchNil"] ORELSE EVERY[resolve_tac ctxt [SwitchCons mode] i,trace_tac ctxt "SwitchCons", HoareRuleTac tac false (i+2), HoareRuleTac tac false (i+1)]; fun while_tac' (t,i) = while_tac ctxt state_kind Inv (HoareRuleTac tac true) tac mode t i; in st |> ( (WlpTac tac i THEN HoareRuleTac tac pre_cond i) ORELSE (FIRST([EVERY[resolve_tac ctxt [Skip mode] i,trace_tac ctxt "Skip"], EVERY[resolve_tac ctxt [BasicCond mode] i, trace_tac ctxt "BasicCond", assertion_simp_tac ctxt state_kind [] i], (resolve_tac ctxt [Basic mode] i THEN trace_tac ctxt "Basic") THEN_MAYBE (assertion_simp_tac ctxt state_kind [] i), (* we don't really need simplificaton here. The question is if it is better to simplify the assertion after each Basic step, so that intermediate assertions stay "small", or if we just accumulate the raw assertions and leave the simplification to the final BasicSimpTac *) EVERY[resolve_tac ctxt [Throw mode] i, trace_tac ctxt "Throw"], (resolve_tac ctxt [Raise mode] i THEN trace_tac ctxt "Raise") THEN_MAYBE (assertion_string_eq_simp_tac ctxt state_kind [] i), cond_tac i, switch_tac i, EVERY[resolve_tac ctxt [Block mode] i, trace_tac ctxt "Block", resolve_tac ctxt [allI] (i+2), resolve_tac ctxt [allI] (i+2), HoareRuleTac tac false (i+2), resolve_tac ctxt [allI] (i+1), in_assertion_simp_tac ctxt state_kind [] (i+1), HoareRuleTac tac false (i+1)], SUBGOAL while_tac' i, SUBGOAL (guard_tac ctxt (annotate_inv orelse strip_guards) (HoareRuleTac tac false) mode) i, EVERY[SUBGOAL (specAnno_tac ctxt state_kind (HoareRuleTac tac true) mode) i], EVERY[SUBGOAL (whileAnnoFix_tac ctxt state_kind (HoareRuleTac tac true) mode) i], EVERY[resolve_tac ctxt [SpecIf mode] i, trace_tac ctxt "SpecIf", assertion_simp_tac ctxt state_kind [] i], (resolve_tac ctxt [Spec mode] i THEN trace_tac ctxt "Spec") THEN_MAYBE (assertion_simp_tac ctxt state_kind [@{thm split_conv}] i), EVERY[resolve_tac ctxt [BindR mode] i, trace_tac ctxt "Bind", resolve_tac ctxt [allI] (i+1), HoareRuleTac tac false (i+1)], EVERY[resolve_tac ctxt [DynCom mode] i, trace_tac ctxt "DynCom"], EVERY[trace_tac ctxt "calling call_tac",SUBGOAL call i], EVERY[trace_tac ctxt "LemmaAnno",SUBGOAL (lemAnno_tac ctxt state_kind mode) i]] @ map (mk_hoare_tac (HoareRuleTac tac) ctxt mode i) hoare_tacs) THEN (if pre_cond then EVERY [trace_tac ctxt "pre_cond", TRY (BasicSimpTac ctxt state_kind true [] tac i), (* FIXME: Do we need TRY *) trace_tac ctxt "after BasicSimpTac"] else (resolve_tac ctxt [subset_refl] i)))) end; in ((K (EVERY [REPEAT (resolve_tac ctxt [allI] 1), HoareRuleTac tac true 1])) THEN_ALL_NEW (prems_tac ctxt)) 1 st (*Procedure specifications may have an locale assumption as premise. These are accumulated by the vcg and are be solved afterward by prems_tac *) end; fun prefer_tac i = (Tactic.defer_tac i THEN PRIMITIVE (Thm.permute_prems 0 ~1)); fun HoareStepTac strip_guards mode state_kind state_space spec_sfx ctxt tac st = let val asms = try (fn () => let val (_,_,_,_,_,G,T,F) = dest_hoare (Logic.strip_assums_concl (Logic.get_goal (Thm.prop_of st) 1)); val params = (strip_vars (Logic.get_goal (Thm.prop_of st) 1)); in mk_proc_assoc (gen_context_thms ctxt mode params G T F) end) () |> the_default []; fun result_tac i = TRY (EVERY [resolve_tac ctxt [Basic mode] i, resolve_tac ctxt [subset_refl] i]); fun call (t,i) = call_tac result_tac mode state_kind state_space ctxt asms spec_sfx t i fun final_simp_tac i = EVERY [full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [mem_Collect_eq]) i, REPEAT (eresolve_tac ctxt [conjE] i), TRY (hyp_subst_tac_thin true ctxt i), BasicSimpTac ctxt state_kind true [] tac i] fun while_annoG_tac (t,i) = whileAnnoG_tac ctxt strip_guards mode t i; in st |> (REPEAT (resolve_tac ctxt [allI] 1) THEN FIRST [resolve_tac ctxt [subset_refl] 1, EVERY[resolve_tac ctxt [Skip mode] 1,TRY (BasicSimpTac ctxt state_kind false [] tac 1)], EVERY[resolve_tac ctxt [BasicCond mode] 1,trace_tac ctxt "BasicCond", TRY (BasicSimpTac ctxt state_kind false [] tac 1)], EVERY[resolve_tac ctxt [Basic mode] 1,TRY (BasicSimpTac ctxt state_kind false [] tac 1)], EVERY[resolve_tac ctxt [Throw mode] 1,TRY (BasicSimpTac ctxt state_kind false [] tac 1)], EVERY[resolve_tac ctxt [Raise mode] 1,TRY (assertion_string_eq_simp_tac ctxt state_kind [] 1)], resolve_tac ctxt [SeqSwap mode] 1 THEN_MAYBE HoareStepTac strip_guards mode state_kind state_space spec_sfx ctxt tac, EVERY[resolve_tac ctxt [BSeq mode] 1, prefer_tac 2 THEN_MAYBE HoareStepTac strip_guards mode state_kind state_space spec_sfx ctxt tac], resolve_tac ctxt [CondSwap mode] 1, resolve_tac ctxt [SwitchNil mode] 1, resolve_tac ctxt [SwitchCons mode] 1, EVERY [SUBGOAL while_annoG_tac 1], EVERY[resolve_tac ctxt [While mode] 1, if mode=Total then wf_tac ctxt 4 else all_tac, BasicSimpTac ctxt state_kind false [] tac 3, if mode=Total then in_rel_simp ctxt 2 THEN (resolve_tac ctxt [allI] 2) else all_tac, BasicSimpTac ctxt state_kind false [] tac 1], resolve_tac ctxt [CatchSwap mode] 1, resolve_tac ctxt [CondCatchSwap mode] 1, EVERY[resolve_tac ctxt [BlockSwap mode] 1, resolve_tac ctxt [allI] 1, resolve_tac ctxt [allI] 1, resolve_tac ctxt [allI] 2, BasicSimpTac ctxt state_kind false [] tac 2], resolve_tac ctxt [GuardsNil mode] 1, resolve_tac ctxt [GuardsConsGuaranteeStrip mode] 1, resolve_tac ctxt [GuardsCons mode] 1, SUBGOAL (guard_tac ctxt strip_guards (fn i => all_tac) mode) 1, EVERY[SUBGOAL (specAnno_tac ctxt state_kind (K all_tac) mode) 1], EVERY[SUBGOAL (whileAnnoFix_tac ctxt state_kind (K all_tac) mode) 1], EVERY[resolve_tac ctxt [SpecIf mode] 1,trace_tac ctxt "SpecIf", TRY (BasicSimpTac ctxt state_kind false [] tac 1)], EVERY[resolve_tac ctxt [Spec mode] 1, TRY (BasicSimpTac ctxt state_kind false [@{thm split_conv}] tac 1)], EVERY[resolve_tac ctxt [BindR mode] 1, resolve_tac ctxt [allI] 2, prefer_tac 2], EVERY[resolve_tac ctxt [FCall mode] 1], EVERY[resolve_tac ctxt [DynCom mode] 1], EVERY[SUBGOAL call 1, BasicSimpTac ctxt state_kind false [] tac 1], EVERY[SUBGOAL (lemAnno_tac ctxt state_kind mode) 1, BasicSimpTac ctxt state_kind false [] tac 1], final_simp_tac 1 ]) end; (*****************************************************************************) (** Generalise verification condition **) (*****************************************************************************) structure RecordSplitState : SPLIT_STATE = struct val globals = @{const_name StateSpace.state.globals}; fun isState (Const _$Abs (s,T,t)) = (case (state_hierarchy T) of ((n,_)::_) => n = "StateSpace.state.state" andalso is_none (try dest_hoare_raw (strip_qnt_body @{const_name All} t)) | _ => false) | isState _ = false; fun isFreeState (Free (_,T)) = (case (state_hierarchy T) of ((n,_)::_) => n = "StateSpace.state.state" | _ => false) | isFreeState _ = false; val abs_state = Option.map snd o first_subterm isFreeState; fun sel_eq (Const (x,_)$_) y = (x=y) | sel_eq t y = raise TERM ("RecordSplitState.sel_eq",[t]); val sel_idx = idx sel_eq; fun bound xs (t as (Const (x,_)$_)) = let val i = sel_idx xs x in if i < 0 then (length xs, xs@[t]) else (i,xs) end | bound xs t = raise TERM ("RecordSplitState.bound",[t]); fun abs_var _ (Const (x,T)$_) = (remdeco' (Long_Name.base_name x),range_type T) | abs_var _ t = raise TERM ("RecordSplitState.abs_var",[t]); fun fld_eq (x, _) y = (x = y) fun fld_idx xs x = idx fld_eq xs x; fun sort_vars ctxt T vars = let val thy = Proof_Context.theory_of ctxt; val (flds,_) = Record.get_recT_fields thy T; val gT = the (AList.lookup (fn (x:string,y) => x=y) flds globals); val (gflds,_) = (Record.get_recT_fields thy gT handle TYPE _ => ([],("",dummyT))); fun compare (Const _$Free _, Const _$(Const _$Free _)) = GREATER | compare (Const (s1,_)$Free _, Const (s2,_)$Free _) = int_ord (fld_idx flds s1, fld_idx flds s2) | compare (Const (s1,_)$(Const _$Free _), Const (s2,_)$(Const _$Free _)) = int_ord (fld_idx gflds s1, fld_idx gflds s2) | compare _ = LESS; in sort (rev_order o compare) vars end; fun fold_state_prop loc glob app abs other inc s res (t as (Const (sel,_)$Free (s',_))) = if s'=s then if is_state_var sel then loc inc res t else raise TERM ("RecordSplitState.fold_state_prop",[t]) else other res t | fold_state_prop loc glob app abs other inc s res (t as ((t1 as (Const (sel,_)))$(t2 as (Const (glb,_)$Free (s',_))))) = if s'=s andalso is_state_var sel andalso (glb=globals) then glob inc res t else let val res1 = fold_state_prop loc glob app abs other inc s res t1 val res2 = fold_state_prop loc glob app abs other inc s res1 t2 in app res1 res2 end | fold_state_prop loc glob app abs other inc s res (t as (Free (s',_))) = if s'=s then raise TERM ("RecordSplitState.fold_state_prop",[t]) else other res t | fold_state_prop loc glob app abs other inc s res (t1$t2) = let val res1 = fold_state_prop loc glob app abs other inc s res t1 val res2 = fold_state_prop loc glob app abs other inc s res1 t2 in app res1 res2 end | fold_state_prop loc glob app abs other inc s res (Abs (x,T,t)) = let val res1 = fold_state_prop loc glob app abs other (inc+1) s res t in abs x T res1 end | fold_state_prop loc glob app abs other inc s res t = other res t fun collect_vars s t = let fun loc _ vars t = snd (bound vars t); fun glob _ vars t = snd (bound vars t); fun app _ vars2 = vars2; fun abs _ _ vars = vars; fun other vars _ = vars; in fold_state_prop loc glob app abs other 0 s [] t end; fun abstract_vars vars s t = let fun loc inc _ t = let val i = fst (bound vars t) in Bound (i+inc) end; fun glob inc _ t = let val i = fst (bound vars t) in Bound (i+inc) end; fun app t1 t2 = t1$t2; fun abs x T t = Abs (x,T,t); fun other _ t = t; val dummy = Bound 0; in fold_state_prop loc glob app abs other 0 s dummy t end; fun split_state ctxt s T t = let val vars = collect_vars s t; val vars' = if Config.get ctxt sort_variables then sort_vars ctxt T vars else vars; in (abstract_vars vars' s t,rev vars') end; fun ex_tac ctxt _ st = Record.split_simp_tac ctxt @{thms simp_thms} (K ~1) 1 st; end; structure FunSplitState : SPLIT_STATE = struct val full_globalsN = @{const_name StateSpace.state.globals}; fun isState (Const _$Abs (s,T,t)) = (case (state_hierarchy T) of ((n,_)::_) => n = "StateSpace.state.state" andalso is_none (try dest_hoare_raw (strip_qnt_body @{const_name All} t)) | _ => false) | isState _ = false; fun isFreeState (Free (_,T)) = (case (state_hierarchy T) of ((n,_)::_) => n = "StateSpace.state.state" | _ => false) | isFreeState _ = false; val abs_state = Option.map snd o first_subterm isFreeState; fun comp_name t = case try (implode o dest_string) t of SOME str => str | NONE => (case t of Free (s,_) => s | Const (s,_) => s | t => raise TERM ("FunSplitState.comp_name",[t])) fun sel_name (Const _$_$name$_) = comp_name name | sel_name t = raise TERM ("FunSplitState.sel_name",[t]); fun sel_raw_name (Const _$_$name$_) = name | sel_raw_name t = raise TERM ("FunSplitState.sel_raw_name",[t]); fun component_type (Const _$_$_$(sel$_)) = range_type (fastype_of sel) | component_type t = raise TERM ("FunSplitState.component_type",[t]); fun component_name (Const _$_$_$((Const (sel,_)$_))) = sel | component_name t = raise TERM ("FunSplitState.component_name",[t]); fun sel_type (Const _$destr$_$_) = range_type (fastype_of destr) | sel_type t = raise TERM ("FunSplitState.sel_type",[t]); fun sel_destr (Const _$destr$_$_) = destr | sel_destr t = raise TERM ("FunSplitState.sel_destr",[t]); fun sel_eq t y = (sel_name t = y) | sel_eq t y = raise TERM ("FunSplitState.sel_eq",[t]); val sel_idx = idx sel_eq; fun bound xs t = let val i = sel_idx xs (sel_name t) in if i < 0 then (length xs, xs@[t]) else (i,xs) end | bound xs t = raise TERM ("FunSplitState.bound",[t]); fun fold_state_prop var app abs other inc s res (t as (Const (@{const_name StateFun.lookup},_)$destr$name$(Const _$Free (s',_)))) = if s'=s then var inc res t else other res t (*raise TERM ("FunSplitState.fold_state_prop",[t])*) | fold_state_prop var app abs other inc s res (t as (Free (s',_))) = if s'=s then raise TERM ("FunSplitState.fold_state_prop",[t]) else other res t | fold_state_prop var app abs other inc s res (t1$t2) = let val res1 = fold_state_prop var app abs other inc s res t1 val res2 = fold_state_prop var app abs other inc s res1 t2 in app res1 res2 end | fold_state_prop var app abs other inc s res (Abs (x,T,t)) = let val res1 = fold_state_prop var app abs other (inc+1) s res t in abs x T res1 end | fold_state_prop var app abs other inc s res t = other res t fun collect_vars s t = let fun var _ vars t = snd (bound vars t); fun app _ vars2 = vars2; fun abs _ _ vars = vars; fun other vars _ = vars; in fold_state_prop var app abs other 0 s [] t end; fun abstract_vars vars s t = let fun var inc _ t = let val i = fst (bound vars t) in Bound (i+inc) end; fun app t1 t2 = t1$t2; fun abs x T t = Abs (x,T,t); fun other _ t = t; val dummy = Bound 0; in fold_state_prop var app abs other 0 s dummy t end; fun sort_vars ctxt vars = let val fld_idx = idx (fn s1:string => fn s2 => s1 = s2); fun compare (_$_$n$(Const (s1,_)$_),_$_$m$(Const (s2,_)$_)) = let val n' = remdeco' (comp_name n); val m' = remdeco' (comp_name m); in if s1 = full_globalsN then if s2 = full_globalsN then string_ord (n',m') else LESS else if s2 = full_globalsN then GREATER else string_ord (n',m') end | compare (t1,t2) = raise TERM ("FunSplitState.sort_vars.compare",[t1,t2]); in sort (rev_order o compare) vars end; fun split_state ctxt s _ t = let val vars = collect_vars s t; val vars' = if Config.get ctxt sort_variables then sort_vars ctxt vars else vars; in (abstract_vars vars' s t,rev vars') end; fun abs_var _ t = (remdeco' (sel_name t), sel_type t); (* Proof for: EX x_1 ... x_n. P x_1 ... x_n * ==> EX s. P (lookup destr_1 "x_1" s) ... (lookup destr_n "x_n" s) * Implementation: * 1. Eliminate existential quantifiers in premise * 2. Instantiate s with: (%x. undefined)("x_1" := constr_1 x_1, ..., "x_n" := constr_n x_n) * 3. Simplify *) local val ss = simpset_of (put_simpset (simpset_of @{theory_context Fun}) @{context} addsimps (@{thm StateFun.lookup_def} :: @{thm StateFun.id_id_cancel} :: @{thms list.inject list.distinct char.inject cong_exp_iff_simps simp_thms}) addsimprocs [Record.simproc, StateFun.lazy_conj_simproc] |> fold Simplifier.add_cong @{thms block_conj_cong}); in fun ex_tac ctxt vs st = let val vs' = rev vs; val (Const (_,exT)$_) = HOLogic.dest_Trueprop (Logic.strip_imp_concl (Logic.get_goal (Thm.prop_of st) 1)); val sT = domain_type (domain_type exT); val s0 = Const (@{const_name HOL.undefined},sT); fun streq (s1:string,s2) = s1=s2 ; fun mk_init [] = [] | mk_init (t::ts) = let val xs = mk_init ts; val n = component_name t; val T = component_type t; in if AList.defined streq xs n then xs else (n,(T,Const (n,sT --> component_type t)$s0))::xs end; fun mk_upd (i,t) xs = let val selN = component_name t; val selT = component_type t; val (_,s) = the (AList.lookup streq xs selN); val strT = domain_type selT; val valT = range_type selT; val constr = destr_to_constr (sel_destr t); val name = (sel_raw_name t); val upd = Const (@{const_name Fun.fun_upd}, (strT --> valT)-->strT-->valT--> (strT --> valT)) $ s $ name $ (constr $ Bound i) in AList.update streq (selN,(selT,upd)) xs end; val upds = fold_index mk_upd vs' (mk_init vs'); val upd = fold (fn (n,(T,upd)) => fn s => Const (n ^ Record.updateN, T --> sT --> sT)$upd$s) upds s0; val inst = fold_rev (Term.abs o (fn t => (sel_name t, sel_type t))) vs upd; fun lift_inst_ex_tac i st = let val rule = Thm.lift_rule (Thm.cprem_of st i) (Drule.incr_indexes st exI); val (_$x) = HOLogic.dest_Trueprop (Logic.strip_assums_concl (hd (Thm.prems_of rule))); val inst_rule = infer_instantiate ctxt [(#1 (dest_Var (head_of x)), Thm.cterm_of ctxt inst)] rule; in (compose_tac ctxt (false,inst_rule, Thm.nprems_of exI) i st) end; in EVERY [REPEAT_DETERM_N (length vs) (eresolve_tac ctxt [exE] 1), lift_inst_ex_tac 1, simp_tac (put_simpset ss ctxt) 1 ] st end end (* Test: What happens when there are no lookups., EX s. True *) end; structure GeneraliseRecord = GeneraliseFun (structure SplitState=RecordSplitState); structure GeneraliseStateFun = GeneraliseFun (structure SplitState=FunSplitState); fun generalise Record = GeneraliseRecord.GENERALISE | generalise Function = GeneraliseStateFun.GENERALISE; (*****************************************************************************) (** record_vanish_tac splits up the records of a verification condition, **) (** trying to generate a predicate without records. **) (** A typical verification condition with a procedure call will have the **) (** form "!!s Z. s=Z ==> ..., where s and Z are records **) (*****************************************************************************) (* FIXME: Check out if removing the useless vars is a performance issue. If so, maybe we can remove all useless vars at once (no iterated simplification) or try to avoid introducing them. Bevore splitting the gaol we can simplifiy the goal with state_simproc this may leed to better performance... *) fun record_vanish_tac ctxt state_kind state_space i = if Config.get ctxt record_vanish then let val rem_useless_vars_simps = [Drule.triv_forall_equality,@{thm Hoare.triv_All_eq},@{thm Hoare.triv_Ex_eq}]; val rem_useless_vars_simpset = empty_simpset ctxt addsimps rem_useless_vars_simps; fun no_spec (t as (Const (@{const_name All},_)$_)) = is_none (try dest_hoare_raw (strip_qnt_body @{const_name All} t)) | no_spec _ = true; fun state_space_no_spec t = if state_space t <> 0 andalso no_spec t then ~1 else 0; in EVERY [trace_tac ctxt "record_vanish_tac -- START --", REPEAT (eresolve_tac ctxt [conjE] i), trace_tac ctxt "record_vanish_tac -- hyp_subst_tac ctxt --", TRY (hyp_subst_tac_thin true ctxt i), full_simp_tac rem_useless_vars_simpset i, (* hyp_subst_tac may have made some state variables unnecessary. We do not want to split them to avoid naming conflicts and increase performance *) trace_tac ctxt "record_vanish_tac -- Splitting records --", if Config.get ctxt use_generalise orelse state_kind = Function then generalise state_kind ctxt i else state_split_simp_tac ctxt rem_useless_vars_simps state_space_no_spec i (*THEN_MAYBE EVERY [trace_tac ctxt "record_vanish_tac -- removing useless vars --", full_simp_tac rem_useless_vars_simpset i, trace_tac ctxt "record_vanish_tac -- STOP --"]*) ] end else all_tac; (* solve_modifies_tac tries to solve modifies-clauses automatically; * The following strategy is followed: * After clar-simplifying the modifies clause we remain with a goal of the form * * EX a b. s(|A := x|) = s(|A:=a,B:=b|) * * (or because of conditional statements conjunctions of these kind of goals) * We split up the state-records and simplify (record_vanish_tac) and get to a goal of the form: * * EX a b. (|A=x,B=B|) = (|A=a,B=b|). * * If the modifies clause was correct we just have to introduce the existential quantifies * and apply reflexivity. * If not we just simplify the goal. *) local val state_fun_update_ss = simpset_of (put_simpset HOL_basic_ss @{context} addsimps ([@{thm StateFun.update_def}, @{thm id_def}, @{thm fun_upd_apply}, @{thm if_True}, @{thm if_False}] @ @{thms list.inject list.distinct char.inject cong_exp_iff_simps simp_thms} @ K_fun_convs) addsimprocs [DistinctTreeProver.distinct_simproc ["distinct_fields", "distinct_fields_globals"]] |> Simplifier.add_cong @{thm imp_cong} (* K_fun_congs FIXME: Stefan fragen*) |> Splitter.add_split @{thm if_split}); in fun solve_modifies_tac ctxt state_kind state_space i st = let val thy = Proof_Context.theory_of ctxt; fun is_split_state (trm as (Const (@{const_name Pure.all},_)$Abs(x,T,t))) = if state_space trm <> 0 then try (fn () => let fun seed (_ $ v $ st) = seed st | seed (_ $ t) = (1,t) (* split only state pair *) | seed t = (~1,t) (* split globals completely *) val all_vars = strip_qnt_vars @{const_name Pure.all} t; val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl t); val ex_vars = strip_qnt_vars @{const_name Ex} concl; val state = Bound (length all_vars + length ex_vars); val (Const (@{const_name HOL.eq},_)$x_upd$x_upd') = strip_qnt_body @{const_name Ex} concl; val (split,sd) = seed x_upd; in if sd = state then split else 0 end) () |> the_default 0 else 0 | is_split_state t = 0; val simp_ctxt = put_simpset HOL_ss ctxt addsimps (@{thm Ex_True} :: @{thm Ex_False} :: Record.get_extinjects thy); fun try_solve Record i = (*(SOLVE*) (((fn k => (TRY (REPEAT_ALL_NEW (resolve_tac ctxt [conjI, impI, allI]) k))) THEN_ALL_NEW (fn k => EVERY [state_split_simp_tac ctxt [] is_split_state k, simp_tac simp_ctxt k THEN_MAYBE rename_goal ctxt remdeco' k ])) i) (*)*) | try_solve _ i = ((fn k => (TRY (REPEAT_ALL_NEW (resolve_tac ctxt [conjI, impI, allI]) k))) THEN_ALL_NEW (fn k => REPEAT (resolve_tac ctxt [exI] k) THEN resolve_tac ctxt [ext] k THEN simp_tac (put_simpset state_fun_update_ss ctxt) k THEN_MAYBE (REPEAT_ALL_NEW (resolve_tac ctxt [conjI,impI,refl]) k))) i in ((trace_tac ctxt "solve_modifies_tac" THEN clarsimp_tac ((ctxt |> put_claset (claset_of @{theory_context HOL}) |> put_simpset (simpset_of @{theory_context Set})) addsimps ([@{thm Hoare.mex_def},@{thm Hoare.meq_def}]@K_convs) addsimprocs (state_upd_simproc Record::(state_simprocs Record)) |> fold Simplifier.add_cong K_congs) i) THEN_MAYBE try_solve state_kind i) st end; end fun proc_lookup_simp_tac ctxt i st = try (fn () => let val name = (Logic.concl_of_goal (Thm.prop_of st) i) |> dest_hoare |> #2 |> strip_comb |> #2 |> last |> strip_comb |> #2 |> last; (* the$(Gamma$name) or the$(strip$Gamma$name) *) val pname = (unsuffix proc_deco (dest_string' name)); val thms = map_filter I (map (try (Proof_Context.get_thm ctxt)) [suffix bodyP pname, suffix (body_def_sfx^"_def") pname, suffix procL pname^"."^ (suffix (body_def_sfx^"_def") pname)]); in simp_tac (put_simpset HOL_basic_ss ctxt addsimps thms @ strip_simps @ @{thms option.sel}) i st end) () |> the_default (Seq.single st); fun proc_lookup_in_dom_simp_tac ctxt i st = try (fn () => let val _$name$_ = (HOLogic.dest_Trueprop (Logic.concl_of_goal (Thm.prop_of st) i)); (* name : Gamma *) val pname = (unsuffix proc_deco (dest_string' name)); val thms = map_filter I (map (try (Proof_Context.get_thm ctxt)) [suffix bodyP pname]); in simp_tac (put_simpset HOL_basic_ss ctxt addsimps (@{thm Hoare.lookup_Some_in_dom} :: @{thm dom_strip} :: thms)) i st end) () |> the_default (Seq.single st); fun HoareRuleTac ctxt insts fixes st = let val annotate_simp_tac = simp_tac (put_simpset HOL_basic_ss ctxt addsimps (anno_defs@normalize_simps) addsimprocs [@{simproc case_prod_beta}]); fun is_com_eq (Const (@{const_name Trueprop},_)$(Const (@{const_name HOL.eq},T)$_$_)) = (case (binder_types T) of (Type (@{type_name Language.com},_)::_) => true | _ => false) | is_com_eq _ = false; fun annotate_tac i st = if is_com_eq (Logic.concl_of_goal (Thm.prop_of st) i) then annotate_simp_tac i st else all_tac st; in ((fn i => REPEAT (resolve_tac ctxt [allI] i)) THEN' Rule_Insts.res_inst_tac ctxt insts fixes st) THEN_ALL_NEW annotate_tac end; fun HoareCallRuleTac state_kind state_space ctxt thms i st = let fun dest_All (Const (@{const_name All},_)$t) = SOME t | dest_All _ = NONE; fun auxvars t = (case (map_filter ((first_subterm is_hoare) o snd) (max_subterms_dest dest_All t)) of ((vars,_)::_) => vars | _ => []); fun auxtype rule = (case (auxvars (Thm.prop_of rule)) of [] => NONE | vs => (case (last vs) of (_,TVar (z,_)) => SOME (z,rule) | _ => NONE)); val thms' = let val auxvs = map fst (auxvars (Logic.concl_of_goal (Thm.prop_of st) i)); val tvar_thms = map_filter auxtype thms in if length thms = length tvar_thms then adapt_aux_var ctxt true auxvs tvar_thms else thms end; val is_sidecondition = not o can dest_hoare; fun solve_sidecondition_tac (t,i) = if is_sidecondition t then FIRST [CHANGED_PROP (wf_tac ctxt i), (*init_conforms_tac state_kind state_space i,*) post_conforms_tac ctxt state_kind i THEN_MAYBE (if is_modifies_clause t then solve_modifies_tac ctxt state_kind state_space i else all_tac), proc_lookup_in_dom_simp_tac ctxt i ] else in_rel_simp ctxt i THEN simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm Un_empty_left},@{thm Un_empty_right}]) i THEN proc_lookup_simp_tac ctxt i fun basic_tac i = (((resolve_tac ctxt thms') THEN_ALL_NEW (fn k => (SUBGOAL solve_sidecondition_tac k))) i) in (basic_tac ORELSE' (fn k => (REPEAT (resolve_tac ctxt [allI] k)) THEN EVERY [resolve_tac ctxt thms' k])) i st end; (* vcg_polish_tac tries to solve modifies-clauses automatically; for other specifications the * records are only splitted and simplified. *) fun vcg_polish_tac solve_modifies ctxt state_kind state_space i = if solve_modifies then solve_modifies_tac ctxt state_kind state_space i else record_vanish_tac ctxt state_kind state_space i THEN_MAYBE EVERY [rename_goal ctxt remdeco' i(*, simp_tac (HOL_basic_ss addsimps @{thms simp_thms})) i*)]; fun is_funtype (Type ("fun", _)) = true | is_funtype _ = false; fun state_kind_of ctxt T = let val thy = Proof_Context.theory_of ctxt; val (s,sT) = nth (fst (Record.get_recT_fields thy T)) 1; in if Long_Name.base_name s = "locals" andalso is_funtype sT then Function else Record end handle Subscript => Record; fun find_state_space_in_triple ctxt t = try (fn () => (case first_subterm is_hoare t of NONE => NONE | SOME (abs_vars,triple) => let val (_,com,_,_,mode,_,_,_) = dest_hoare_raw triple; val T = fastype_of1 (map snd abs_vars,com) val Type(_,state_spaceT::_) = T; val SOME Tids = stateT_ids state_spaceT; in SOME (Tids,mode, state_kind_of ctxt state_spaceT) end)) () |> Option.join; fun get_state_space_in_subset_eq ctxt t = (* get state type from the following kind of terms: P <= Q, s: P *) try (fn () => let val (subset_eq,_) = (strip_comb o HOLogic.dest_Trueprop o strip_qnt_body @{const_name Pure.all}) t; val Ts = map snd (strip_vars t); val T = fastype_of1 (Ts,subset_eq); val Type (_, [_,Type (_, [Type (_, [state_spaceT]), _])]) = T; (* also works for "in": x : P *) val SOME Tids = stateT_ids state_spaceT; in (Tids,Partial, state_kind_of ctxt state_spaceT) end) (); fun get_state_space ctxt i st = (case try (Logic.concl_of_goal (Thm.prop_of st)) i of SOME t => (case find_state_space_in_triple ctxt t of SOME sp => SOME sp | NONE => get_state_space_in_subset_eq ctxt t) | NONE => NONE); fun mk_hoare_tac hoare_tac finish_tac annotate_inv exnames strip_guards spec_sfx ctxt i st = case get_state_space ctxt i st of SOME (Tids,mode,kind) => SELECT_GOAL (hoare_tac annotate_inv exnames strip_guards mode kind (is_state_space_var Tids) spec_sfx ctxt (finish_tac kind (is_state_space_var Tids))) i st | NONE => no_tac st fun vcg_tac spec_sfx strip_guards exnames ctxt i st = mk_hoare_tac HoareTac (vcg_polish_tac (spec_sfx="_modifies") ctxt) (spec_sfx="_modifies") exnames (strip_guards="true") spec_sfx ctxt i st; fun hoare_tac spec_sfx strip_guards _ ctxt i st = let fun tac state_kind state_space i = if spec_sfx="_modifies" then solve_modifies_tac ctxt state_kind state_space i else all_tac; in mk_hoare_tac HoareTac tac (spec_sfx="_modifies") [] (strip_guards="true") spec_sfx ctxt i st end; fun hoare_raw_tac spec_sfx strip_guards exnames ctxt i st = mk_hoare_tac HoareTac (K (K (K all_tac))) (spec_sfx="_modifies") [] (strip_guards="true") spec_sfx ctxt i st; fun hoare_step_tac spec_sfx strip_guards exnames ctxt i st = mk_hoare_tac (K (K HoareStepTac)) (vcg_polish_tac (spec_sfx="_modifies") ctxt) false [] (strip_guards="true") spec_sfx ctxt i st; fun hoare_rule_tac ctxt thms i st = SUBGOAL (fn _ => (case get_state_space ctxt i st of SOME (Tids,_,kind) => HoareCallRuleTac kind (is_state_space_var Tids) ctxt thms i | NONE => error "could not find proper state space type (structure or record) in goal")) i st; (*** Methods ***) val hoare_rule = Rule_Insts.method HoareRuleTac hoare_rule_tac; val argP = Args.name --| @{keyword "="} -- Args.name val argsP = Scan.repeat argP val default_args = [("spec","spec"),("strip_guards","false")] val vcg_simp_modifiers = [Args.add -- Args.colon >> K (Method.modifier vcg_simp_add \<^here>), Args.del -- Args.colon >> K (Method.modifier vcg_simp_del \<^here>)]; fun assocs2 key = map snd o filter (curry (op =) key o fst); fun gen_simp_method tac = Scan.lift (argsP >> (fn args => args @ default_args)) --| Method.sections vcg_simp_modifiers >> (fn args => fn ctxt => Method.SIMPLE_METHOD' (tac ("_" ^ the (AList.lookup (op =) args "spec")) (the (AList.lookup (op =) args "strip_guards")) (assocs2 "exspec" args) ctxt)); val hoare = gen_simp_method hoare_tac; val hoare_raw = gen_simp_method hoare_raw_tac; val vcg = gen_simp_method vcg_tac; val vcg_step = gen_simp_method hoare_step_tac; val trace_hoare_users = Unsynchronized.ref false fun print_subgoal_tac ctxt s i = SUBGOAL (fn (prem, _) => trace_tac ctxt (s ^ (Syntax.string_of_term ctxt prem))) i fun mk_hoare_thm thm _ ctxt _ i = EVERY [resolve_tac ctxt [thm] i, if !trace_hoare_users then print_subgoal_tac ctxt "Tracing: " i else all_tac] val vcg_hoare_add = Thm.declaration_attribute (fn thm => add_hoare_tacs [(Thm.derivation_name thm, mk_hoare_thm thm)]) exception UNDEF val vcg_hoare_del = Thm.declaration_attribute (fn _ => fn _ => raise UNDEF) (* setup theory *) val _ = Theory.setup (Attrib.setup @{binding vcg_simp} (Attrib.add_del vcg_simp_add vcg_simp_del) "declaration of Simplifier rule for vcg" #> Attrib.setup @{binding vcg_hoare} (Attrib.add_del vcg_hoare_add vcg_hoare_del) "declaration of wp rule for vcg") (*#> add_wp_tacs initial_wp_tacs*) end; diff --git a/thys/Stellar_Quorums/ROOT b/thys/Stellar_Quorums/ROOT --- a/thys/Stellar_Quorums/ROOT +++ b/thys/Stellar_Quorums/ROOT @@ -1,9 +1,9 @@ chapter AFP session Stellar_Quorums (AFP) = HOL + - options [timeout = 600] + options [timeout = 600, document=pdf] theories Stellar_Quorums document_files "root.tex" "root.bib" diff --git a/thys/Stirling_Formula/Stirling_Formula.thy b/thys/Stirling_Formula/Stirling_Formula.thy --- a/thys/Stirling_Formula/Stirling_Formula.thy +++ b/thys/Stirling_Formula/Stirling_Formula.thy @@ -1,723 +1,722 @@ (* File: Stirling_Formula.thy Author: Manuel Eberl A proof of Stirling's formula, i.e. an asymptotic approximation of the Gamma function and the factorial. *) section \Stirling's Formula\ theory Stirling_Formula imports "HOL-Analysis.Analysis" "Landau_Symbols.Landau_More" begin context begin text \ First, we define the $S_n^*$ from Jameson's article: \ private definition S' :: "nat \ real \ real" where "S' n x = 1/(2*x) + (\r=1.. Next, the trapezium (also called $T$ in Jameson's article): \ private definition T :: "real \ real" where "T x = 1/(2*x) + 1/(2*(x+1))" text \ Now we define The difference $\Delta(x)$: \ private definition D :: "real \ real" where "D x = T x - ln (x + 1) + ln x" private lemma S'_telescope_trapezium: assumes "n > 0" shows "S' n x = (\rrrrrrrr=0..r=0..r=Suc 0.. = (\r=1.. + ?a + ?b = S' n x" by (simp add: S'_def Suc) finally show ?thesis .. qed (insert assms, simp_all) private lemma stirling_trapezium: assumes x: "(x::real) > 0" shows "D x \ {0 .. 1/(12*x^2) - 1/(12 * (x+1)^2)}" proof - define y where "y = 1 / (2*x + 1)" from x have y: "y > 0" "y < 1" by (simp_all add: divide_simps y_def) from x have "D x = T x - ln ((x + 1) / x)" by (subst ln_div) (simp_all add: D_def) also from x have "(x + 1) / x = 1 + 1 / x" by (simp add: field_simps) finally have D: "D x = T x - ln (1 + 1/x)" . from y have "(\n. y * y^n) sums (y * (1 / (1 - y)))" by (intro geometric_sums sums_mult) simp_all hence "(\n. y ^ Suc n) sums (y / (1 - y))" by simp also from x have "y / (1 - y) = 1 / (2*x)" by (simp add: y_def divide_simps) finally have *: "(\n. y ^ Suc n) sums (1 / (2*x))" . from y have "(\n. (-y) * (-y)^n) sums ((-y) * (1 / (1 - (-y))))" by (intro geometric_sums sums_mult) simp_all hence "(\n. (-y) ^ Suc n) sums (-(y / (1 + y)))" by simp also from x have "y / (1 + y) = 1 / (2*(x+1))" by (simp add: y_def divide_simps) finally have **: "(\n. (-y) ^ Suc n) sums (-(1 / (2*(x+1))))" . from sums_diff[OF * **] have sum1: "(\n. y ^ Suc n - (-y) ^ Suc n) sums T x" by (simp add: T_def) from y have "abs y < 1" "abs (-y) < 1" by simp_all from sums_diff[OF this[THEN ln_series']] have "(\n. y ^ n / real n - (-y) ^ n / real n) sums (ln (1 + y) - ln (1 - y))" by simp also from y have "ln (1 + y) - ln (1 - y) = ln ((1 + y) / (1 - y))" by (simp add: ln_div) also from x have "(1 + y) / (1 - y) = 1 + 1/x" by (simp add: divide_simps y_def) finally have "(\n. y^n / real n - (-y)^n / real n) sums ln (1 + 1/x)" . hence sum2: "(\n. y^Suc n / real (Suc n) - (-y)^Suc n / real (Suc n)) sums ln (1 + 1/x)" by (subst sums_Suc_iff) simp from sum2 sum1 have "ln (1 + 1/x) \ T x" proof (rule sums_le [OF allI, rotated]) fix n :: nat show "y ^ Suc n / real (Suc n) - (-y) ^ Suc n / real (Suc n) \ y^Suc n - (-y) ^ Suc n" proof (cases "even n") case True hence eq: "A - (-y) ^ Suc n / B = A + y ^ Suc n / B" "A - (-y) ^ Suc n = A + y ^ Suc n" for A B by simp_all from y show ?thesis unfolding eq by (intro add_mono) (auto simp: divide_simps) qed simp_all qed hence "D x \ 0" by (simp add: D) define c where "c = (\n. if even n then 2 * (1 - 1 / real (Suc n)) else 0)" note sums_diff[OF sum1 sum2] also have "(\n. y ^ Suc n - (-y) ^ Suc n - (y ^ Suc n / real (Suc n) - (-y) ^ Suc n / real (Suc n))) = (\n. c n * y ^ Suc n)" by (intro ext) (simp add: c_def algebra_simps) finally have sum3: "(\n. c n * y ^ Suc n) sums D x" by (simp add: D) from y have "(\n. y^2 * (of_nat (Suc n) * y^n)) sums (y^2 * (1 / (1 - y)^2))" by (intro sums_mult geometric_deriv_sums) simp_all hence "(\n. of_nat (Suc n) * y^(n+2)) sums (y^2 / (1 - y)^2)" by (simp add: algebra_simps power2_eq_square) also from x have "y^2 / (1 - y)^2 = 1 / (4*x^2)" by (simp add: y_def divide_simps) finally have *: "(\n. real (Suc n) * y ^ (Suc (Suc n))) sums (1 / (4 * x\<^sup>2))" by simp from y have "(\n. y^2 * (of_nat (Suc n) * (-y)^n)) sums (y^2 * (1 / (1 - -y)^2))" by (intro sums_mult geometric_deriv_sums) simp_all hence "(\n. of_nat (Suc n) * (-y)^(n+2)) sums (y^2 / (1 + y)^2)" by (simp add: algebra_simps power2_eq_square) also from x have "y^2 / (1 + y)^2 = 1 / (2^2*(x+1)^2)" unfolding power_mult_distrib [symmetric] by (simp add: y_def divide_simps add_ac) finally have **: "(\n. real (Suc n) * (- y) ^ (Suc (Suc n))) sums (1 / (4 * (x + 1)\<^sup>2))" by simp define d where "d = (\n. if even n then 2 * real n else 0)" note sums_diff[OF * **] also have "(\n. real (Suc n) * y^(Suc (Suc n)) - real (Suc n) * (-y)^(Suc (Suc n))) = (\n. d (Suc n) * y ^ Suc (Suc n))" by (intro ext) (simp_all add: d_def) finally have "(\n. d n * y ^ Suc n) sums (1 / (4 * x\<^sup>2) - 1 / (4 * (x + 1)\<^sup>2))" by (subst (asm) sums_Suc_iff) (simp add: d_def) from sums_mult[OF this, of "1/3"] x have sum4: "(\n. d n / 3 * y ^ Suc n) sums (1 / (12 * x^2) - 1 / (12 * (x + 1)^2))" by (simp add: field_simps) have "D x \ (1 / (12 * x^2) - 1 / (12 * (x + 1)^2))" proof (intro sums_le [OF _ sum3 sum4] allI) fix n :: nat define c' :: "nat \ real" where "c' = (\n. if odd n \ n = 0 then 0 else if n = 2 then 4/3 else 2)" show "c n * y ^ Suc n \ d n / 3 * y ^ Suc n" proof (intro mult_right_mono) have "c n \ c' n" by (simp add: c_def c'_def) also consider "n = 0" | "n = 1" | "n = 2" | "n \ 3" by force hence "c' n \ d n / 3" by cases (simp_all add: c'_def d_def) finally show "c n \ d n / 3" . qed (insert y, simp) qed with \D x \ 0\ show ?thesis by simp qed text \ The following functions correspond to the $p_n(x)$ from the article. The special case $n = 0$ would not, strictly speaking, be necessary, but it allows some theorems to work even without the precondition $n \neq 0$: \ private definition p :: "nat \ real \ real" where "p n x = (if n = 0 then 1/x else (\r We can write the Digamma function in terms of @{term S'}: \ private lemma S'_LIMSEQ_Digamma: assumes x: "x \ 0" shows "(\n. ln (real n) - S' n x - 1/(2*x)) \ Digamma x" proof - define c where "c = (\n. ln (real n) - (\rn. 1 / (2 * (x + real n)) = c n - (ln (real n) - S' n x - 1/(2*x))) at_top" using eventually_gt_at_top[of "0::nat"] proof eventually_elim fix n :: nat assume n: "n > 0" have "c n - (ln (real n) - S' n x - 1/(2*x)) = -(\rr=1..r=1..rn. 1 / (2 * (x + real n))) \ 0" by (rule real_tendsto_divide_at_top tendsto_const filterlim_tendsto_pos_mult_at_top filterlim_tendsto_add_at_top filterlim_real_sequentially | simp)+ ultimately have "(\n. c n - (ln (real n) - S' n x - 1/(2*x))) \ 0" by (blast intro: Lim_transform_eventually) from tendsto_minus[OF this] have "(\n. (ln (real n) - S' n x - 1/(2*x)) - c n) \ 0" by simp moreover from Digamma_LIMSEQ[OF x] have "c \ Digamma x" by (simp add: c_def) ultimately show "(\n. ln (real n) - S' n x - 1/(2*x)) \ Digamma x" by (rule Lim_transform [rotated]) qed text \ Moreover, we can give an expansion of @{term S'} with the @{term p} as variation terms. \ private lemma S'_approx: "S' n x = ln (real n + x) - ln x + p n x" proof (cases "n = 0") case True thus ?thesis by (simp add: p_def S'_def) next case False hence "S' n x = (\r = (\r = (\rr We define the limit of the @{term p} (simply called $p(x)$ in Jameson's article): \ private definition P :: "real \ real" where "P x = (\n. D (real n + x))" private lemma D_summable: assumes x: "x > 0" shows "summable (\n. D (real n + x))" proof - have *: "summable (\n. 1 / (12 * (x + real n)\<^sup>2) - 1 / (12 * (x + real (Suc n))\<^sup>2))" by (rule telescope_summable' real_tendsto_divide_at_top tendsto_const filterlim_tendsto_pos_mult_at_top filterlim_pow_at_top filterlim_tendsto_add_at_top filterlim_real_sequentially | simp)+ show "summable (\n. D (real n + x))" proof (rule summable_comparison_test[OF _ *], rule exI[of _ 2], safe) fix n :: nat assume "n \ 2" show "norm (D (real n + x)) \ 1 / (12 * (x + real n)^2) - 1 / (12 * (x + real (Suc n))^2)" using stirling_trapezium[of "real n + x"] x by (auto simp: algebra_simps) qed qed private lemma p_LIMSEQ: assumes x: "x > 0" shows "(\n. p n x) \ P x" proof (rule Lim_transform_eventually) from D_summable[OF x] have "(\n. D (real n + x)) sums P x" unfolding P_def by (simp add: sums_iff) then show "(\n. \r P x" by (simp add: sums_def) moreover from eventually_gt_at_top[of 1] show "eventually (\n. (\r This gives us an expansion of the Digamma function: \ lemma Digamma_approx: assumes x: "(x :: real) > 0" shows "Digamma x = ln x - 1 / (2 * x) - P x" proof - have "eventually (\n. ln (inverse (1 + x / real n)) + ln x - 1/(2*x) - p n x = ln (real n) - S' n x - 1/(2*x)) at_top" using eventually_gt_at_top[of "1::nat"] proof eventually_elim fix n :: nat assume n: "n > 1" have "ln (real n) - S' n x = ln ((real n) / (real n + x)) + ln x - p n x" using assms n unfolding S'_approx by (subst ln_div) (auto simp: algebra_simps) also from n have "real n / (real n + x) = inverse (1 + x / real n)" by (simp add: field_simps) finally show "ln (inverse (1 + x / real n)) + ln x - 1/(2*x) - p n x = ln (real n) - S' n x - 1/(2*x)" by simp qed moreover have "(\n. ln (inverse (1 + x / real n)) + ln x - 1/(2*x) - p n x) \ ln (inverse (1 + 0)) + ln x - 1/(2*x) - P x" by (rule tendsto_intros p_LIMSEQ x real_tendsto_divide_at_top filterlim_real_sequentially | simp)+ hence "(\n. ln (inverse (1 + x / real n)) + ln x - 1/(2*x) - p n x) \ ln x - 1/(2*x) - P x" by simp ultimately have "(\n. ln (real n) - S' n x - 1 / (2 * x)) \ ln x - 1/(2*x) - P x" by (blast intro: Lim_transform_eventually) moreover from x have "(\n. ln (real n) - S' n x - 1 / (2 * x)) \ Digamma x" by (intro S'_LIMSEQ_Digamma) simp_all ultimately show "Digamma x = ln x - 1 / (2 * x) - P x" by (rule LIMSEQ_unique [rotated]) qed text \ Next, we derive some bounds on @{term "P"}: \ private lemma p_ge_0: "x > 0 \ p n x \ 0" using stirling_trapezium[of "real n + x" for n] by (auto simp add: p_def intro!: sum_nonneg) private lemma P_ge_0: "x > 0 \ P x \ 0" by (rule tendsto_lowerbound[OF p_LIMSEQ]) (insert p_ge_0[of x], simp_all) private lemma p_upper_bound: assumes "x > 0" "n > 0" shows "p n x \ 1/(12*x^2)" proof - from assms have "p n x = (\r \ (\r = 1 / (12 * x\<^sup>2) - 1 / (12 * (real n + x)\<^sup>2)" by (subst sum_lessThan_telescope') simp also have "\ \ 1 / (12 * x^2)" by simp finally show ?thesis . qed private lemma P_upper_bound: assumes "x > 0" shows "P x \ 1/(12*x^2)" proof (rule tendsto_upperbound) show "eventually (\n. p n x \ 1 / (12 * x^2)) at_top" using eventually_gt_at_top[of 0] by eventually_elim (use p_upper_bound[of x] assms in auto) show "(\n. p n x) \ P x" by (simp add: assms p_LIMSEQ) qed auto text \ We can now use this approximation of the Digamma function to approximate @{term ln_Gamma} (since the former is the derivative of the latter). We therefore define the function $g$ from Jameson's article, which measures the difference between @{term ln_Gamma} and its approximation: \ private definition g :: "real \ real" where "g x = ln_Gamma x - (x - 1/2) * ln x + x" private lemma DERIV_g: "x > 0 \ (g has_field_derivative -P x) (at x)" unfolding g_def [abs_def] using Digamma_approx[of x] by (auto intro!: derivative_eq_intros simp: field_simps) private lemma isCont_P: assumes "x > 0" shows "isCont P x" proof - define D' :: "real \ real" where "D' = (\x. - 1 / (2 * x^2 * (x+1)^2))" have DERIV_D: "(D has_field_derivative D' x) (at x)" if "x > 0" for x unfolding D_def [abs_def] D'_def T_def by (insert that, (rule derivative_eq_intros refl | simp)+) (simp add: power2_eq_square divide_simps, (simp add: algebra_simps)?) note this [THEN DERIV_chain2, derivative_intros] have "(P has_field_derivative (\n. D' (real n + x))) (at x)" unfolding P_def [abs_def] proof (rule has_field_derivative_series') show "convex {x/2<..}" by simp next fix n :: nat and y :: real assume y: "y \ {x/2<..}" with assms have "y > 0" by simp thus "((\a. D (real n + a)) has_real_derivative D' (real n + y)) (at y within {x/2<..})" by (auto intro!: derivative_eq_intros) next from assms D_summable[of x] show "summable (\n. D (real n + x))" by simp next show "uniformly_convergent_on {x/2<..} (\n x. \i {x/2<..}" with assms have "y > 0" by auto have "norm (D' (real n + y)) = (1 / (2 * (y + real n)^2)) * (1 / (y + real (Suc n))^2)" by (simp add: D'_def add_ac) also from y assms have "\ \ (1 / (2 * (x/2)^2)) * (1 / (real (Suc n))^2)" by (intro mult_mono divide_left_mono power_mono) simp_all also have "1 / (real (Suc n))^2 = inverse ((real (Suc n))^2)" by (simp add: field_simps) finally show "norm (D' (real n + y)) \ (1 / (2 * (x/2)^2)) * inverse ((real (Suc n))^2)" . next show "summable (\n. (1 / (2 * (x/2)^2)) * inverse ((real (Suc n))^2))" by (subst summable_Suc_iff, intro summable_mult inverse_power_summable) simp_all qed qed (insert assms, simp_all add: interior_open) thus ?thesis by (rule DERIV_isCont) qed private lemma P_continuous_on [THEN continuous_on_subset]: "continuous_on {0<..} P" by (intro continuous_at_imp_continuous_on ballI isCont_P) auto private lemma P_integrable: assumes a: "a > 0" shows "P integrable_on {a..}" proof - define f where "f = (\n x. if x \ {a..real n} then P x else 0)" show "P integrable_on {a..}" proof (rule dominated_convergence) fix n :: nat from a have "P integrable_on {a..real n}" by (intro integrable_continuous_real P_continuous_on) auto hence "f n integrable_on {a..real n}" by (rule integrable_eq) (simp add: f_def) thus "f n integrable_on {a..}" by (rule integrable_on_superset) (auto simp: f_def) next fix n :: nat show "norm (f n x) \ of_real (1/12) * (1 / x^2)" if "x \ {a..}" for x using a P_ge_0 P_upper_bound by (auto simp: f_def) next show "(\x::real. of_real (1/12) * (1 / x^2)) integrable_on {a..}" using has_integral_inverse_power_to_inf[of 2 a] a by (intro integrable_on_cmult_left) auto next show "(\n. f n x) \ P x" if "x\{a..}" for x proof - have "eventually (\n. real n \ x) at_top" using filterlim_real_sequentially by (simp add: filterlim_at_top) with that have "eventually (\n. f n x = P x) at_top" by (auto elim!: eventually_mono simp: f_def) thus "(\n. f n x) \ P x" by (simp add: tendsto_eventually) qed qed qed private definition c :: real where "c = integral {1..} (\x. -P x) + g 1" text \ We can now give bounds on @{term g}: \ private lemma g_bounds: assumes x: "x \ 1" shows "g x \ {c..c + 1/(12*x)}" proof - from assms have int_nonneg: "integral {x..} P \ 0" by (intro Henstock_Kurzweil_Integration.integral_nonneg P_integrable) (auto simp: P_ge_0) have int_upper_bound: "integral {x..} P \ 1/(12*x)" proof (rule has_integral_le) from x show "(P has_integral integral {x..} P) {x..}" by (intro integrable_integral P_integrable) simp_all from x has_integral_mult_right[OF has_integral_inverse_power_to_inf[of 2 x], of "1/12"] show "((\x. 1/(12*x^2)) has_integral (1/(12*x))) {x..}" by (simp add: field_simps) qed (insert P_upper_bound x, simp_all) note DERIV_g [THEN DERIV_chain2, derivative_intros] from assms have int1: "((\x. -P x) has_integral (g x - g 1)) {1..x}" by (intro fundamental_theorem_of_calculus) (auto simp: has_field_derivative_iff_has_vector_derivative [symmetric] intro!: derivative_eq_intros) from x have int2: "((\x. -P x) has_integral integral {x..} (\x. -P x)) {x..}" by (intro integrable_integral integrable_neg P_integrable) simp_all from has_integral_Un[OF int1 int2] x have "((\x. - P x) has_integral g x - g 1 - integral {x..} P) ({1..x} \ {x..})" by (simp add: max_def) also from x have "{1..x} \ {x..} = {1..}" by auto finally have "((\x. -P x) has_integral g x - g 1 - integral {x..} P) {1..}" . moreover have "((\x. -P x) has_integral integral {1..} (\x. -P x)) {1..}" by (intro integrable_integral integrable_neg P_integrable) simp_all ultimately have "g x - g 1 - integral {x..} P = integral {1..} (\x. -P x)" by (simp add: has_integral_unique) hence "g x = c + integral {x..} P" by (simp add: c_def algebra_simps) with int_upper_bound int_nonneg show "g x \ {c..c + 1/(12*x)}" by simp qed text \ Finally, we have bounds on @{term ln_Gamma}, @{term Gamma}, and @{term fact}. \ private lemma ln_Gamma_bounds_aux: "x \ 1 \ ln_Gamma x \ c + (x - 1/2) * ln x - x" "x \ 1 \ ln_Gamma x \ c + (x - 1/2) * ln x - x + 1/(12*x)" using g_bounds[of x] by (simp_all add: g_def) private lemma Gamma_bounds_aux: assumes x: "x \ 1" shows "Gamma x \ exp c * x powr (x - 1/2) / exp x" "Gamma x \ exp c * x powr (x - 1/2) / exp x * exp (1/(12*x))" proof - have "exp (ln_Gamma x) \ exp (c + (x - 1/2) * ln x - x)" by (subst exp_le_cancel_iff, rule ln_Gamma_bounds_aux) (simp add: x) with x show "Gamma x \ exp c * x powr (x - 1/2) / exp x" by (simp add: Gamma_real_pos_exp exp_add exp_diff powr_def del: exp_le_cancel_iff) next have "exp (ln_Gamma x) \ exp (c + (x - 1/2) * ln x - x + 1/(12*x))" by (subst exp_le_cancel_iff, rule ln_Gamma_bounds_aux) (simp add: x) with x show "Gamma x \ exp c * x powr (x - 1/2) / exp x * exp (1/(12*x))" by (simp add: Gamma_real_pos_exp exp_add exp_diff powr_def del: exp_le_cancel_iff) qed private lemma Gamma_asymp_equiv_aux: "Gamma \[at_top] (\x. exp c * x powr (x - 1/2) / exp x)" proof (rule asymp_equiv_sandwich) include asymp_equiv_notation show "eventually (\x. exp c * x powr (x - 1/2) / exp x \ Gamma x) at_top" "eventually (\x. exp c * x powr (x - 1/2) / exp x * exp (1/(12*x)) \ Gamma x) at_top" using eventually_ge_at_top[of "1::real"] by (eventually_elim; use Gamma_bounds_aux in force)+ have "((\x::real. exp (1 / (12 * x))) \ exp 0) at_top" by (rule tendsto_intros real_tendsto_divide_at_top filterlim_tendsto_pos_mult_at_top)+ (simp_all add: filterlim_ident) hence "(\x. exp (1 / (12 * x))) \ (\x. 1 :: real)" by (intro asymp_equivI') simp_all hence "(\x. exp c * x powr (x - 1 / 2) / exp x * 1) \ (\x. exp c * x powr (x - 1 / 2) / exp x * exp (1 / (12 * x)))" by (intro asymp_equiv_mult asymp_equiv_refl) (simp add: asymp_equiv_sym) thus "(\x. exp c * x powr (x - 1 / 2) / exp x) \ (\x. exp c * x powr (x - 1 / 2) / exp x * exp (1 / (12 * x)))" by simp qed simp_all private lemma exp_1_powr_real [simp]: "exp (1::real) powr x = exp x" by (simp add: powr_def) private lemma fact_asymp_equiv_aux: "fact \[at_top] (\x. exp c * sqrt (real x) * (real x / exp 1) powr real x)" proof - include asymp_equiv_notation have "fact \ (\n. Gamma (real (Suc n)))" by (simp add: Gamma_fact) also have "eventually (\n. Gamma (real (Suc n)) = real n * Gamma (real n)) at_top" using eventually_gt_at_top[of "0::nat"] by eventually_elim (insert Gamma_plus1[of "real n" for n], auto simp: add_ac of_nat_in_nonpos_Ints_iff) also have "(\n. Gamma (real n)) \ (\n. exp c * real n powr (real n - 1/2) / exp (real n))" by (rule asymp_equiv_compose'[OF Gamma_asymp_equiv_aux] filterlim_real_sequentially)+ also have "eventually (\n. real n * (exp c * real n powr (real n - 1 / 2) / exp (real n)) = exp c * sqrt (real n) * (real n / exp 1) powr real n) at_top" using eventually_gt_at_top[of "0::nat"] proof eventually_elim fix n :: nat assume n: "n > 0" thus "real n * (exp c * real n powr (real n - 1 / 2) / exp (real n)) = exp c * sqrt (real n) * (real n / exp 1) powr real n" by (subst powr_diff) (simp_all add: powr_divide powr_half_sqrt field_simps) qed finally show ?thesis by - (simp_all add: asymp_equiv_mult) qed text \ We still need to determine the constant term @{term c}, which we do using Wallis' product formula: $$\prod_{n=1}^\infty \frac{4n^2}{4n^2-1} = \frac{\pi}{2}$$ \ private lemma powr_mult_2: "(x::real) > 0 \ x powr (y * 2) = (x^2) powr y" by (subst mult.commute, subst powr_powr [symmetric]) (simp add: powr_numeral) private lemma exp_mult_2: "exp (y * 2 :: real) = exp y * exp y" by (subst exp_add [symmetric]) simp private lemma exp_c: "exp c = sqrt (2*pi)" proof - include asymp_equiv_notation define p where "p = (\n. \k=1..n. (4*real k^2) / (4*real k^2 - 1))" have p_0 [simp]: "p 0 = 1" by (simp add: p_def) have p_Suc: "p (Suc n) = p n * (4 * real (Suc n)^2) / (4 * real (Suc n)^2 - 1)" for n unfolding p_def by (subst prod.nat_ivl_Suc') simp_all have p: "p = (\n. 16 ^ n * fact n ^ 4 / (fact (2 * n))\<^sup>2 / (2 * real n + 1))" proof fix n :: nat have "p n = (\k=1..n. (2*real k)^2 / (2*real k - 1) / (2 * real k + 1))" unfolding p_def by (intro prod.cong refl) (simp add: field_simps power2_eq_square) also have "\ = (\k=1..n. (2*real k)^2 / (2*real k - 1)) / (\k=1..n. (2 * real (Suc k) - 1))" by (simp add: prod_dividef prod.distrib add_ac) also have "(\k=1..n. (2 * real (Suc k) - 1)) = (\k=Suc 1..Suc n. (2 * real k - 1))" by (subst prod.atLeast_Suc_atMost_Suc_shift) simp_all also have "\ = (\k=1..n. (2 * real k - 1)) * (2 * real n + 1)" by (induction n) (simp_all add: prod.nat_ivl_Suc') also have "(\k = 1..n. (2 * real k)\<^sup>2 / (2 * real k - 1)) / \ = (\k = 1..n. (2 * real k)^2 / (2 * real k - 1)^2) / (2 * real n + 1)" unfolding power2_eq_square by (simp add: prod.distrib prod_dividef) also have "(\k = 1..n. (2 * real k)^2 / (2 * real k - 1)^2) = (\k = 1..n. (2 * real k)^4 / ((2*real k)*(2 * real k - 1))^2)" - by (intro prod.cong refl) - (simp add: divide_simps, (simp add: field_simps power2_eq_square eval_nat_numeral)) + by (rule prod.cong) (simp_all add: power2_eq_square eval_nat_numeral) also have "\ = 16^n * fact n^4 / (\k=1..n. (2*real k) * (2*real k - 1))^2" by (simp add: prod.distrib prod_dividef fact_prod prod_power_distrib [symmetric] prod_constant) also have "(\k=1..n. (2*real k) * (2*real k - 1)) = fact (2*n)" by (induction n) (simp_all add: algebra_simps prod.nat_ivl_Suc') finally show "p n = 16 ^ n * fact n ^ 4 / (fact (2 * n))\<^sup>2 / (2 * real n + 1)" . qed have "p \ (\n. 16 ^ n * fact n ^ 4 / (fact (2 * n))\<^sup>2 / (2 * real n + 1))" by (simp add: p) also have "\ \ (\n. 16^n * (exp c * sqrt (real n) * (real n / exp 1) powr real n)^4 / (exp c * sqrt (real (2*n)) * (real (2*n) / exp 1) powr real (2*n))^2 / (2 * real n + 1))" (is "_ \ ?f") by (intro asymp_equiv_mult asymp_equiv_divide asymp_equiv_refl mult_nat_left_at_top fact_asymp_equiv_aux asymp_equiv_power asymp_equiv_compose'[OF fact_asymp_equiv_aux]) simp_all also have "eventually (\n. \ n = exp c ^ 2 / (4 + 2/n)) at_top" using eventually_gt_at_top[of "0::nat"] proof eventually_elim fix n :: nat assume n: "n > 0" have [simp]: "16^n = 4^n * (4^n :: real)" by (simp add: power_mult_distrib [symmetric]) from n have "?f n = exp c ^ 2 * (n / (2*(2*n+1)))" by (simp add: power_mult_distrib divide_simps powr_mult real_sqrt_power_even) (simp add: field_simps power2_eq_square eval_nat_numeral powr_mult_2 exp_mult_2 powr_realpow) also from n have "\ = exp c ^ 2 / (4 + 2/n)" by (simp add: field_simps) finally show "?f n = \" . qed also have "(\x. 4 + 2 / real x) \ (\x. 4)" by (subst asymp_equiv_add_right) auto finally have "p \ exp c ^ 2 / 4" by (rule asymp_equivD_const) (simp_all add: asymp_equiv_divide) moreover have "p \ pi / 2" unfolding p_def by (rule wallis) ultimately have "exp c ^ 2 / 4 = pi / 2" by (rule LIMSEQ_unique) hence "2 * pi = exp c ^ 2" by simp also have "sqrt (exp c ^ 2) = exp c" by simp finally show "exp c = sqrt (2 * pi)" .. qed private lemma c: "c = ln (2*pi) / 2" proof - note exp_c [symmetric] also have "ln (exp c) = c" by simp finally show ?thesis by (simp add: ln_sqrt) qed text \ This gives us the final bounds: \ theorem Gamma_bounds: assumes "x \ 1" shows "Gamma x \ sqrt (2*pi/x) * (x / exp 1) powr x" (is ?th1) "Gamma x \ sqrt (2*pi/x) * (x / exp 1) powr x * exp (1 / (12 * x))" (is ?th2) proof - from assms have "exp c * x powr (x - 1/2) / exp x = sqrt (2*pi/x) * (x / exp 1) powr x" by (subst powr_diff) (simp add: exp_c real_sqrt_divide powr_divide powr_half_sqrt field_simps) with Gamma_bounds_aux[OF assms] show ?th1 ?th2 by simp_all qed theorem ln_Gamma_bounds: assumes "x \ 1" shows "ln_Gamma x \ ln (2*pi/x) / 2 + x * ln x - x" (is ?th1) "ln_Gamma x \ ln (2*pi/x) / 2 + x * ln x - x + 1/(12*x)" (is ?th2) proof - from ln_Gamma_bounds_aux[OF assms] assms show ?th1 ?th2 by (simp_all add: c field_simps ln_div) from assms have "exp c * x powr (x - 1/2) / exp x = sqrt (2*pi/x) * (x / exp 1) powr x" by (subst powr_diff) (simp add: exp_c real_sqrt_divide powr_divide powr_half_sqrt field_simps) qed theorem fact_bounds: assumes "n > 0" shows "(fact n :: real) \ sqrt (2*pi*n) * (n / exp 1) ^ n" (is ?th1) "(fact n :: real) \ sqrt (2*pi*n) * (n / exp 1) ^ n * exp (1 / (12 * n))" (is ?th2) proof - from assms have n: "real n \ 1" by simp from assms Gamma_plus1[of "real n"] have "real n * Gamma (real n) = Gamma (real (Suc n))" by (simp add: of_nat_in_nonpos_Ints_iff add_ac) also have "Gamma (real (Suc n)) = fact n" by (subst Gamma_fact [symmetric]) simp finally have *: "fact n = real n * Gamma (real n)" by simp have "2*pi/n = 2*pi*n / n^2" by (simp add: power2_eq_square) also have "sqrt \ = sqrt (2*pi*n) / n" by (subst real_sqrt_divide) simp_all also have "real n * \ = sqrt (2*pi*n)" by simp finally have **: "real n * sqrt (2*pi/real n) = sqrt (2*pi*real n)" . note * also note Gamma_bounds(2)[OF n] also have "real n * (sqrt (2 * pi / real n) * (real n / exp 1) powr real n * exp (1 / (12 * real n))) = (real n * sqrt (2*pi/n)) * (n / exp 1) powr n * exp (1 / (12 * n))" by (simp add: algebra_simps) also from n have "(real n / exp 1) powr real n = (real n / exp 1) ^ n" by (subst powr_realpow) simp_all also note ** finally show ?th2 by - (insert assms, simp_all) have "sqrt (2*pi*n) * (n / exp 1) powr n = n * (sqrt (2*pi/n) * (n / exp 1) powr n)" by (subst ** [symmetric]) (simp add: field_simps) also from assms have "\ \ real n * Gamma (real n)" by (intro mult_left_mono Gamma_bounds(1)) simp_all also from n have "(real n / exp 1) powr real n = (real n / exp 1) ^ n" by (subst powr_realpow) simp_all also note * [symmetric] finally show ?th1 . qed theorem ln_fact_bounds: assumes "n > 0" shows "ln (fact n :: real) \ ln (2*pi*n)/2 + n * ln n - n" (is ?th1) "ln (fact n :: real) \ ln (2*pi*n)/2 + n * ln n - n + 1/(12*real n)" (is ?th2) proof - have "ln (fact n :: real) \ ln (sqrt (2*pi*real n)*(real n/exp 1)^n)" using fact_bounds(1)[OF assms] assms by (subst ln_le_cancel_iff) auto also from assms have "ln (sqrt (2*pi*real n)*(real n/exp 1)^n) = ln (2*pi*n)/2 + n * ln n - n" by (simp add: ln_mult ln_div ln_realpow ln_sqrt field_simps) finally show ?th1 . next have "ln (fact n :: real) \ ln (sqrt (2*pi*real n) * (real n/exp 1)^n * exp (1/(12*real n)))" using fact_bounds(2)[OF assms] assms by (subst ln_le_cancel_iff) auto also from assms have "\ = ln (2*pi*n)/2 + n * ln n - n + 1/(12*real n)" by (simp add: ln_mult ln_div ln_realpow ln_sqrt field_simps) finally show ?th2 . qed theorem Gamma_asymp_equiv: "Gamma \[at_top] (\x. sqrt (2*pi/x) * (x / exp 1) powr x :: real)" proof - note Gamma_asymp_equiv_aux also have "eventually (\x. exp c * x powr (x - 1/2) / exp x = sqrt (2*pi/x) * (x / exp 1) powr x) at_top" using eventually_gt_at_top[of "0::real"] proof eventually_elim fix x :: real assume x: "x > 0" thus "exp c * x powr (x - 1/2) / exp x = sqrt (2*pi/x) * (x / exp 1) powr x" by (subst powr_diff) (simp add: exp_c powr_half_sqrt powr_divide field_simps real_sqrt_divide) qed finally show ?thesis . qed theorem fact_asymp_equiv: "fact \[at_top] (\n. sqrt (2*pi*n) * (n / exp 1) ^ n :: real)" proof - note fact_asymp_equiv_aux also have "eventually (\n. exp c * sqrt (real n) = sqrt (2 * pi * real n)) at_top" using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: exp_c real_sqrt_mult) also have "eventually (\n. (n / exp 1) powr n = (n / exp 1) ^ n) at_top" using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: powr_realpow) finally show ?thesis . qed end end diff --git a/thys/Taylor_Models/Interval_Approximation.thy b/thys/Taylor_Models/Interval_Approximation.thy --- a/thys/Taylor_Models/Interval_Approximation.thy +++ b/thys/Taylor_Models/Interval_Approximation.thy @@ -1,774 +1,774 @@ section \Approximate Operations on Intervals of Floating Point Numbers\ theory Interval_Approximation imports "HOL-Decision_Procs.Approximation_Bounds" Interval begin lifting_update float.lifting \ \TODO: in Float!\ lifting_forget float.lifting text \TODO: many of the lemmas should move to theories Float or Approximation (the latter should be based on type @{type interval}.\ subsection "Intervals with Floating Point Bounds" lift_definition round_interval :: "nat \ float interval \ float interval" is "\p. \(l, u). (float_round_down p l, float_round_up p u)" by (auto simp: intro!: float_round_down_le float_round_up_le) lemma lower_round_ivl[simp]: "lower (round_interval p x) = float_round_down p (lower x)" by transfer auto lemma upper_round_ivl[simp]: "upper (round_interval p x) = float_round_up p (upper x)" by transfer auto lemma round_ivl_correct: "set_of A \ set_of (round_interval prec A)" by (auto simp: set_of_eq float_round_down_le float_round_up_le) lift_definition truncate_ivl :: "nat \ real interval \ real interval" is "\p. \(l, u). (truncate_down p l, truncate_up p u)" by (auto intro!: truncate_down_le truncate_up_le) lemma lower_truncate_ivl[simp]: "lower (truncate_ivl p x) = truncate_down p (lower x)" by transfer auto lemma upper_truncate_ivl[simp]: "upper (truncate_ivl p x) = truncate_up p (upper x)" by transfer auto lemma truncate_ivl_correct: "set_of A \ set_of (truncate_ivl prec A)" by (auto simp: set_of_eq intro!: truncate_down_le truncate_up_le) lift_definition real_interval::"float interval \ real interval" is "\(l, u). (real_of_float l, real_of_float u)" by auto lemma lower_real_interval[simp]: "lower (real_interval x) = lower x" by transfer auto lemma upper_real_interval[simp]: "upper (real_interval x) = upper x" by transfer auto definition "set_of' x = (case x of None \ UNIV | Some i \ set_of (real_interval i))" lemma real_interval_min_interval[simp]: "real_interval (min_interval a b) = min_interval (real_interval a) (real_interval b)" by (auto simp: interval_eq_set_of_iff set_of_eq real_of_float_min) lemma real_interval_max_interval[simp]: "real_interval (max_interval a b) = max_interval (real_interval a) (real_interval b)" by (auto simp: interval_eq_set_of_iff set_of_eq real_of_float_max) subsection \Intervals for standard functions\ lift_definition power_float_interval :: "nat \ nat \ float interval \ float interval" is "\p n (l, u). float_power_bnds p n l u" using float_power_bnds by (auto simp: bnds_power dest!: float_power_bnds[OF sym]) lemma lower_power_float_interval[simp]: "lower (power_float_interval p n x) = fst (float_power_bnds p n (lower x) (upper x))" by transfer auto lemma upper_power_float_interval[simp]: "upper (power_float_interval p n x) = snd (float_power_bnds p n (lower x) (upper x))" by transfer auto lemma power_float_intervalI: "x \\<^sub>i real_interval X \ x ^ n \\<^sub>i real_interval (power_float_interval p n X)" using float_power_bnds[OF prod.collapse] by (auto simp: set_of_eq ) lift_definition mult_float_interval::"nat \ float interval \ float interval \ float interval" is "\prec. \(a1, a2). \(b1, b2). bnds_mult prec a1 a2 b1 b2" by (auto dest!: bnds_mult[OF sym]) lemma lower_mult_float_interval[simp]: "lower (mult_float_interval p x y) = fst (bnds_mult p (lower x) (upper x) (lower y) (upper y))" by transfer auto lemma upper_mult_float_interval[simp]: "upper (mult_float_interval p x y) = snd (bnds_mult p (lower x) (upper x) (lower y) (upper y))" by transfer auto lemma mult_float_interval: "set_of (real_interval A) * set_of (real_interval B) \ set_of (real_interval (mult_float_interval prec A B))" proof - let ?bm = "bnds_mult prec (lower A) (upper A) (lower B) (upper B)" show ?thesis using bnds_mult[of "fst ?bm" "snd ?bm", simplified, OF refl] by (auto simp: set_of_eq set_times_def) qed lemma mult_float_intervalI: "x * y \\<^sub>i (real_interval (mult_float_interval prec A B))" if "x \\<^sub>i real_interval A" "y \\<^sub>i real_interval B" using mult_float_interval[of A B] that by (auto simp: ) lift_definition sqrt_float_interval::"nat \ float interval \ float interval" is "\prec. \(lx, ux). (lb_sqrt prec lx, ub_sqrt prec ux)" using bnds_sqrt' by auto (meson order_trans real_sqrt_le_iff) lemma lower_float_interval[simp]: "lower (sqrt_float_interval prec X) = lb_sqrt prec (lower X)" by transfer auto lemma upper_float_interval[simp]: "upper (sqrt_float_interval prec X) = ub_sqrt prec (upper X)" by transfer auto lemma sqrt_float_interval: "sqrt ` set_of (real_interval X) \ set_of (real_interval (sqrt_float_interval prec X))" using bnds_sqrt by (auto simp: set_of_eq) lemma sqrt_float_intervalI: "sqrt x \\<^sub>i real_interval (sqrt_float_interval p X)" if "x \ set_of (real_interval X)" using sqrt_float_interval[of X p] that by auto lemmas [simp del] = lb_arctan.simps ub_arctan.simps lemma lb_arctan: "arctan (real_of_float x) \ y \ real_of_float (lb_arctan prec x) \ y" and ub_arctan: "y \ arctan x \ y \ ub_arctan prec x" for x::float and y::real using arctan_boundaries[of x prec] by auto lift_definition arctan_float_interval :: "nat \ float interval \ float interval" is "\prec. \(lx, ux). (lb_arctan prec lx, ub_arctan prec ux)" by (auto intro!: lb_arctan ub_arctan arctan_monotone') lemma lower_arctan_float_interval[simp]: "lower (arctan_float_interval p x) = lb_arctan p (lower x)" by transfer auto lemma upper_arctan_float_interval[simp]: "upper (arctan_float_interval p x) = ub_arctan p (upper x)" by transfer auto lemma arctan_float_interval: "arctan ` set_of (real_interval x) \ set_of (real_interval (arctan_float_interval p x))" by (auto simp: set_of_eq intro!: lb_arctan ub_arctan arctan_monotone') lemma arctan_float_intervalI: "arctan x \\<^sub>i real_interval (arctan_float_interval p X)" if "x \ set_of (real_interval X)" using arctan_float_interval[of X p] that by auto lemma bnds_cos_lower: "\x. real_of_float xl \ x \ x \ real_of_float xu \ cos x \ y \ real_of_float (fst (bnds_cos prec xl xu)) \ y" and bnds_cos_upper: "\x. real_of_float xl \ x \ x \ real_of_float xu \ y \ cos x \ y \ real_of_float (snd (bnds_cos prec xl xu))" for xl xu::float and y::real using bnds_cos[of "fst (bnds_cos prec xl xu)" "snd (bnds_cos prec xl xu)" prec] by force+ lift_definition cos_float_interval :: "nat \ float interval \ float interval" is "\prec. \(lx, ux). bnds_cos prec lx ux" using bnds_cos by auto (metis (full_types) order_refl order_trans) lemma lower_cos_float_interval[simp]: "lower (cos_float_interval p x) = fst (bnds_cos p (lower x) (upper x))" by transfer auto lemma upper_cos_float_interval[simp]: "upper (cos_float_interval p x) = snd (bnds_cos p (lower x) (upper x))" by transfer auto lemma cos_float_interval: "cos ` set_of (real_interval x) \ set_of (real_interval (cos_float_interval p x))" by (auto simp: set_of_eq bnds_cos_lower bnds_cos_upper) lemma cos_float_intervalI: "cos x \\<^sub>i real_interval (cos_float_interval p X)" if "x \ set_of (real_interval X)" using cos_float_interval[of X p] that by auto lemma lb_exp: "exp x \ y \ lb_exp prec x \ y" and ub_exp: "y \ exp x \ y \ ub_exp prec x" for x::float and y::real using exp_boundaries[of x prec] by auto lift_definition exp_float_interval :: "nat \ float interval \ float interval" is "\prec. \(lx, ux). (lb_exp prec lx, ub_exp prec ux)" by (auto simp: lb_exp ub_exp) lemma lower_exp_float_interval[simp]: "lower (exp_float_interval p x) = lb_exp p (lower x)" by transfer auto lemma upper_exp_float_interval[simp]: "upper (exp_float_interval p x) = ub_exp p (upper x)" by transfer auto lemma exp_float_interval: "exp ` set_of (real_interval x) \ set_of (real_interval (exp_float_interval p x))" using exp_boundaries apply (auto simp: set_of_eq) apply (smt exp_le_cancel_iff) apply (smt exp_le_cancel_iff) done lemma exp_float_intervalI: "exp x \\<^sub>i real_interval (exp_float_interval p X)" if "x \ set_of (real_interval X)" using exp_float_interval[of X p] that by auto lemmas [simp del] = lb_ln.simps ub_ln.simps lemma lb_lnD: "y \ ln x \ 0 < real_of_float x" if "lb_ln prec x = Some y" using lb_ln[OF that[symmetric]] by auto lemma ub_lnD: "ln x \ y\ 0 < real_of_float x" if "ub_ln prec x = Some y" using ub_ln[OF that[symmetric]] by auto lift_definition(code_dt) ln_float_interval :: "nat \ float interval \ float interval option" is "\prec. \(lx, ux). Option.bind (lb_ln prec lx) (\l. Option.bind (ub_ln prec ux) (\u. Some (l, u)))" by (auto simp: pred_option_def bind_eq_Some_conv ln_le_cancel_iff[symmetric] simp del: ln_le_cancel_iff dest!: lb_lnD ub_lnD) lemma ln_float_interval_eq_Some_conv[simp]: "ln_float_interval p x = Some y \ lb_ln p (lower x) = Some (lower y) \ ub_ln p (upper x) = Some (upper y)" by transfer (auto simp: bind_eq_Some_conv) lemma ln_float_interval: "ln ` set_of (real_interval x) \ set_of (real_interval y)" if "ln_float_interval p x = Some y" using that by (simp add: set_of_eq) (smt atLeastAtMost_iff bnds_ln image_subset_iff) lemma ln_float_intervalI: "ln x \ set_of' (ln_float_interval p X)" if "x \\<^sub>i (real_interval X)" using ln_float_interval[of p X] that by (auto simp: set_of'_def split: option.splits) lift_definition(code_dt) powr_float_interval :: "nat \ float interval \ float interval \ float interval option" is "\prec. \(l1, u1). \(l2, u2). bnds_powr prec l1 u1 l2 u2" by (auto simp: pred_option_def dest!: bnds_powr[OF sym]) lemma powr_float_interval: "{x powr y | x y. x \ set_of (real_interval X) \ y \ set_of (real_interval Y)} \ set_of (real_interval R)" if "powr_float_interval prec X Y = Some R" using that by transfer (auto dest!: bnds_powr[OF sym]) lemma powr_float_intervalI: "x powr y \ set_of' (powr_float_interval p X Y)" if "x \\<^sub>i real_interval X" "y \\<^sub>i real_interval Y" using powr_float_interval[of p X Y] that by (auto simp: set_of'_def split: option.splits) lift_definition(code_dt) inverse_float_interval::"nat \ float interval \ float interval option" is "\prec (l, u). if (0 < l \ u < 0) then Some (float_divl prec 1 u, float_divr prec 1 l) else None" by (auto intro!: order_trans[OF float_divl] order_trans[OF _ float_divr] simp: divide_simps) lemma inverse_float_interval_eq_Some_conv[simp]: defines "one \ (1::float)" shows "inverse_float_interval p X = Some R \ (lower X > 0 \ upper X < 0) \ lower R = float_divl p one (upper X) \ upper R = float_divr p one (lower X)" by clarsimp (transfer fixing: one, force simp: one_def split: if_splits) lemma inverse_float_interval: "inverse ` set_of (real_interval X) \ set_of (real_interval Y)" if "inverse_float_interval p X = Some Y" using that apply (clarsimp simp: set_of_eq) by (intro order_trans[OF float_divl] order_trans[OF _ float_divr] conjI) (auto simp: divide_simps) lemma inverse_float_intervalI: "x \ set_of (real_interval X) \ inverse x \ set_of' (inverse_float_interval p X)" using inverse_float_interval[of p X] by (auto simp: set_of'_def split: option.splits) lift_definition pi_float_interval::"nat \ float interval" is "\prec. (lb_pi prec, ub_pi prec)" using pi_boundaries by (auto intro: order_trans) lemma lower_pi_float_interval[simp]: "lower (pi_float_interval prec) = lb_pi prec" by transfer auto lemma upper_pi_float_interval[simp]: "upper (pi_float_interval prec) = ub_pi prec" by transfer auto lemma pi_float_interval: "pi \ set_of (real_interval (pi_float_interval prec))" using pi_boundaries by (auto simp: set_of_eq) lemma real_interval_abs_interval[simp]: "real_interval (abs_interval x) = abs_interval (real_interval x)" by (auto simp: interval_eq_set_of_iff set_of_eq real_of_float_max real_of_float_min) lift_definition floor_float_interval::"float interval \ float interval" is "\(l, u). (floor_fl l, floor_fl u)" by (auto intro!: floor_mono simp: floor_fl.rep_eq) lemma lower_floor_float_interval[simp]: "lower (floor_float_interval x) = floor_fl (lower x)" by transfer auto lemma upper_floor_float_interval[simp]: "upper (floor_float_interval x) = floor_fl (upper x)" by transfer auto lemma floor_float_intervalI: "\x\ \\<^sub>i real_interval (floor_float_interval X)" if "x \\<^sub>i real_interval X" using that by (auto simp: set_of_eq floor_fl_def floor_mono) lemma in_intervalI: "x \\<^sub>i X" if "lower X \ x" "x \ upper X" using that by (auto simp: set_of_eq) abbreviation in_real_interval ("(_/ \\<^sub>r _)" [51, 51] 50) where "x \\<^sub>r X \ x \\<^sub>i real_interval X" lemma in_real_intervalI: "x \\<^sub>r X" if "lower X \ x" "x \ upper X" for x::real and X::"float interval" using that by (intro in_intervalI) auto lemma lower_Interval: "lower (Interval x) = fst x" and upper_Interval: "upper (Interval x) = snd x" if "fst x \ snd x" using that by (auto simp: lower_def upper_def Interval_inverse split_beta') definition all_in_i :: "'a::preorder list \ 'a interval list \ bool" (infix "(all'_in\<^sub>i)" 50) where "x all_in\<^sub>i I = (length x = length I \ (\i < length I. x ! i \\<^sub>i I ! i))" definition all_in :: "real list \ float interval list \ bool" (infix "(all'_in)" 50) where "x all_in I = (length x = length I \ (\i < length I. x ! i \\<^sub>r I ! i))" definition all_subset :: "'a::order interval list \ 'a interval list \ bool" (infix "(all'_subset)" 50) where "I all_subset J = (length I = length J \ (\i < length I. set_of (I!i) \ set_of (J!i)))" lemmas [simp] = all_in_def all_subset_def lemma all_subsetD: assumes "I all_subset J" assumes "x all_in I" shows "x all_in J" using assms by (auto simp: set_of_eq; fastforce) lemma plus_down_mono: "plus_down p a b \ plus_down p c d" if "a + b \ c + d" by (auto simp: plus_down_def intro!: truncate_down_mono that) lemma plus_up_mono: "plus_up p a b \ plus_up p c d" if "a + b \ c + d" by (auto simp: plus_up_def intro!: truncate_up_mono that) lemma round_interval_mono: "set_of (round_interval prec X) \ set_of (round_interval prec Y)" if "set_of X \ set_of Y" using that by transfer (auto simp: float_round_down.rep_eq float_round_up.rep_eq truncate_down_mono truncate_up_mono) lemma mult_mono_nonpos_nonneg: "a * b \ c * d" if "a \ c" "a \ 0" "0 \ d" "d \ b" for a b c d::"'a::ordered_ring" apply (rule order_trans[OF mult_left_mono_neg[OF \d \ b\]]) subgoal using that by auto by (rule mult_right_mono; fact) lemma mult_mono_nonneg_nonpos: "b * a \ d * c" if "a \ c" "c \ 0" "0 \ d" "d \ b" for a b c d::"'a::ordered_ring" apply (rule order_trans[OF mult_right_mono_neg[OF \d \ b\]]) subgoal using that by auto by (rule mult_left_mono; fact) lemma mult_mono_nonpos_nonpos: "a * b \ c * d" if "a \ c" "a \ 0" "b \ d" "d \ 0" for a b c d::real apply (rule order_trans[OF mult_left_mono_neg[OF \d \ b\]]) subgoal using that by auto by (rule mult_right_mono_neg; fact) lemma mult_float_mono1: notes mono_rules = plus_down_mono add_mono nprt_mono nprt_le_zero zero_le_pprt pprt_mono shows "a \ b \ ab \ bb \ aa \ a \ b \ ba \ ac \ ab \ bb \ bc \ plus_down prec (nprt aa * pprt bc) (plus_down prec (nprt ba * nprt bc) (plus_down prec (pprt aa * pprt ac) (pprt ba * nprt ac))) \ plus_down prec (nprt a * pprt bb) (plus_down prec (nprt b * nprt bb) (plus_down prec (pprt a * pprt ab) (pprt b * nprt ab)))" apply (rule order_trans) apply (rule mono_rules | assumption)+ apply (rule mult_mono_nonpos_nonneg) apply (rule mono_rules | assumption)+ apply (rule mult_mono_nonpos_nonpos) apply (rule mono_rules | assumption)+ apply (rule mult_mono) apply (rule mono_rules | assumption)+ apply (rule mult_mono_nonneg_nonpos) apply (rule mono_rules | assumption)+ by (rule order_refl)+ lemma mult_float_mono2: notes mono_rules = plus_up_mono add_mono nprt_mono nprt_le_zero zero_le_pprt pprt_mono shows "a \ b \ ab \ bb \ aa \ a \ b \ ba \ ac \ ab \ bb \ bc \ plus_up prec (pprt b * pprt bb) (plus_up prec (pprt a * nprt bb) (plus_up prec (nprt b * pprt ab) (nprt a * nprt ab))) \ plus_up prec (pprt ba * pprt bc) (plus_up prec (pprt aa * nprt bc) (plus_up prec (nprt ba * pprt ac) (nprt aa * nprt ac)))" apply (rule order_trans) apply (rule mono_rules | assumption)+ apply (rule mult_mono) apply (rule mono_rules | assumption)+ apply (rule mult_mono_nonneg_nonpos) apply (rule mono_rules | assumption)+ apply (rule mult_mono_nonpos_nonneg) apply (rule mono_rules | assumption)+ apply (rule mult_mono_nonpos_nonpos) apply (rule mono_rules | assumption)+ by (rule order_refl)+ lemma mult_float_interval_mono: "set_of (mult_float_interval prec A B) \ set_of (mult_float_interval prec X Y)" if "set_of A \ set_of X" "set_of B \ set_of Y" using that apply transfer unfolding bnds_mult_def atLeastatMost_subset_iff float_plus_down.rep_eq float_plus_up.rep_eq by (auto simp: float_plus_down.rep_eq float_plus_up.rep_eq mult_float_mono1 mult_float_mono2) lemma Ivl_simps[simp]: "lower (Ivl a b) = min a b" "upper (Ivl a b) = b" subgoal by transfer simp subgoal by transfer simp done lemmas [simp del] = power_down.simps(2) power_up.simps(2) lemmas power_down_simp = power_down.simps(2) lemmas power_up_simp = power_up.simps(2) lemma power_down_even_nonneg: "even n \ 0 \ power_down p x n" by (induct p x n rule: power_down.induct) (auto simp: power_down_simp simp del: odd_Suc_div_two intro!: truncate_down_nonneg ) lemma truncate_down_less_zero_iff[simp]: "truncate_down p x < 0 \ x < 0" by (metis le_less_trans not_less_iff_gr_or_eq truncate_down truncate_down_pos truncate_down_zero) lemma truncate_down_nonneg_iff[simp]: "truncate_down p x \ 0 \ x \ 0" using truncate_down_less_zero_iff[of p x] truncate_down_nonneg[of x p] by linarith lemma truncate_down_eq_zero_iff[simp]: "truncate_down prec x = 0 \ x = 0" by (metis not_less_iff_gr_or_eq truncate_down_less_zero_iff truncate_down_pos truncate_down_zero) lemma power_down_eq_zero_iff[simp]: "power_down prec b n = 0 \ b = 0 \ n \ 0" proof (induction n arbitrary: b rule: less_induct) case (less x) then show ?case using power_down_simp[of _ _ "x - 1"] by (cases x) (auto simp add: div2_less_self) qed lemma power_down_nonneg_iff[simp]: "power_down prec b n \ 0 \ even n \ b \ 0" proof (induction n arbitrary: b rule: less_induct) case (less x) show ?case using less(1)[of "x - 1" b] power_down_simp[of _ _ "x - 1"] - by (cases x) (auto simp: sign_simps zero_le_mult_iff) + by (cases x) (auto simp: algebra_split_simps zero_le_mult_iff) qed lemma power_down_neg_iff[simp]: "power_down prec b n < 0 \ b < 0 \ odd n" using power_down_nonneg_iff[of prec b n] by (auto simp del: power_down_nonneg_iff) lemma power_down_nonpos_iff[simp]: notes [simp del] = power_down_neg_iff power_down_eq_zero_iff shows "power_down prec b n \ 0 \ b < 0 \ odd n \ b = 0 \ n \ 0" using power_down_neg_iff[of prec b n] power_down_eq_zero_iff[of prec b n] by auto lemma power_down_mono: "power_down prec a n \ power_down prec b n" if "((0 \ a \ a \ b)\(odd n \ a \ b) \ (even n \ a \ 0 \ b \ a))" using that proof (induction n arbitrary: a b rule: less_induct) case (less i) show ?case proof (cases i) case j: (Suc j) note IH = less[unfolded j even_Suc not_not] note [simp del] = power_down.simps show ?thesis proof cases assume [simp]: "even j" have "a * power_down prec a j \ b * power_down prec b j" by (smt IH(1) IH(2) \even j\ lessI mult_mono' mult_mono_nonpos_nonneg power_down_even_nonneg) then have "truncate_down (Suc prec) (a * power_down prec a j) \ truncate_down (Suc prec) (b * power_down prec b j)" by (auto intro!: truncate_down_mono simp: abs_le_square_iff[symmetric] abs_real_def) then show ?thesis unfolding j by (simp add: power_down_simp) next assume [simp]: "odd j" have "power_down prec 0 (Suc (j div 2)) \ - power_down prec b (Suc (j div 2))" if "b < 0" "even (j div 2)" apply (rule order_trans[where y=0]) using IH that by (auto simp: div2_less_self) then have "truncate_down (Suc prec) ((power_down prec a (Suc (j div 2)))\<^sup>2) \ truncate_down (Suc prec) ((power_down prec b (Suc (j div 2)))\<^sup>2)" using IH by (auto intro!: truncate_down_mono intro: order_trans[where y=0] simp: abs_le_square_iff[symmetric] abs_real_def div2_less_self) then show ?thesis unfolding j by (simp add: power_down_simp) qed qed simp qed lemma truncate_up_nonneg: "0 \ truncate_up p x" if "0 \ x" by (simp add: that truncate_up_le) lemma truncate_up_pos: "0 < truncate_up p x" if "0 < x" by (meson less_le_trans that truncate_up) lemma truncate_up_less_zero_iff[simp]: "truncate_up p x < 0 \ x < 0" proof - have f1: "\n r. truncate_up n r + truncate_down n (- 1 * r) = 0" by (simp add: truncate_down_uminus_eq) have f2: "(\v0 v1. truncate_up v0 v1 + truncate_down v0 (- 1 * v1) = 0) = (\v0 v1. truncate_up v0 v1 = - 1 * truncate_down v0 (- 1 * v1))" by (auto simp: truncate_up_eq_truncate_down) have f3: "\x1. ((0::real) < x1) = (\ x1 \ 0)" by fastforce have "(- 1 * x \ 0) = (0 \ x)" by force then have "0 \ x \ \ truncate_down p (- 1 * x) \ 0" using f3 by (meson truncate_down_pos) then have "(0 \ truncate_up p x) \ (\ 0 \ x)" using f2 f1 truncate_up_nonneg by force then show ?thesis by linarith qed lemma truncate_up_nonneg_iff[simp]: "truncate_up p x \ 0 \ x \ 0" using truncate_up_less_zero_iff[of p x] truncate_up_nonneg[of x] by linarith lemma power_up_even_nonneg: "even n \ 0 \ power_up p x n" by (induct p x n rule: power_up.induct) (auto simp: power_up.simps simp del: odd_Suc_div_two intro!: ) lemma truncate_up_eq_zero_iff[simp]: "truncate_up prec x = 0 \ x = 0" by (metis not_less_iff_gr_or_eq truncate_up_less_zero_iff truncate_up_pos truncate_up_zero) lemma power_up_eq_zero_iff[simp]: "power_up prec b n = 0 \ b = 0 \ n \ 0" proof (induction n arbitrary: b rule: less_induct) case (less x) then show ?case using power_up_simp[of _ _ "x - 1"] - by (cases x) (auto simp: sign_simps zero_le_mult_iff div2_less_self) + by (cases x) (auto simp: algebra_split_simps zero_le_mult_iff div2_less_self) qed lemma power_up_nonneg_iff[simp]: "power_up prec b n \ 0 \ even n \ b \ 0" proof (induction n arbitrary: b rule: less_induct) case (less x) show ?case using less(1)[of "x - 1" b] power_up_simp[of _ _ "x - 1"] - by (cases x) (auto simp: sign_simps zero_le_mult_iff) + by (cases x) (auto simp: algebra_split_simps zero_le_mult_iff) qed lemma power_up_neg_iff[simp]: "power_up prec b n < 0 \ b < 0 \ odd n" using power_up_nonneg_iff[of prec b n] by (auto simp del: power_up_nonneg_iff) lemma power_up_nonpos_iff[simp]: notes [simp del] = power_up_neg_iff power_up_eq_zero_iff shows "power_up prec b n \ 0 \ b < 0 \ odd n \ b = 0 \ n \ 0" using power_up_neg_iff[of prec b n] power_up_eq_zero_iff[of prec b n] by auto lemma power_up_mono: "power_up prec a n \ power_up prec b n" if "((0 \ a \ a \ b)\(odd n \ a \ b) \ (even n \ a \ 0 \ b \ a))" using that proof (induction n arbitrary: a b rule: less_induct) case (less i) show ?case proof (cases i) case j: (Suc j) note IH = less[unfolded j even_Suc not_not] note [simp del] = power_up.simps show ?thesis proof cases assume [simp]: "even j" have "a * power_up prec a j \ b * power_up prec b j" by (smt IH(1) IH(2) \even j\ lessI mult_mono' mult_mono_nonpos_nonneg power_up_even_nonneg) then have "truncate_up prec (a * power_up prec a j) \ truncate_up prec (b * power_up prec b j)" by (auto intro!: truncate_up_mono simp: abs_le_square_iff[symmetric] abs_real_def) then show ?thesis unfolding j by (simp add: power_up_simp) next assume [simp]: "odd j" have "power_up prec 0 (Suc (j div 2)) \ - power_up prec b (Suc (j div 2))" if "b < 0" "even (j div 2)" apply (rule order_trans[where y=0]) using IH that by (auto simp: div2_less_self) then have "truncate_up prec ((power_up prec a (Suc (j div 2)))\<^sup>2) \ truncate_up prec ((power_up prec b (Suc (j div 2)))\<^sup>2)" using IH by (auto intro!: truncate_up_mono intro: order_trans[where y=0] simp: abs_le_square_iff[symmetric] abs_real_def div2_less_self) then show ?thesis unfolding j by (simp add: power_up_simp) qed qed simp qed lemma set_of_subset_iff: "set_of X \ set_of Y \ lower Y \ lower X \ upper X \ upper Y" for X Y::"'a::linorder interval" by (auto simp: set_of_eq subset_iff) lemma power_float_interval_mono: "set_of (power_float_interval prec n A) \ set_of (power_float_interval prec n B)" if "set_of A \ set_of B" proof - define la where "la = real_of_float (lower A)" define ua where "ua = real_of_float (upper A)" define lb where "lb = real_of_float (lower B)" define ub where "ub = real_of_float (upper B)" have ineqs: "lb \ la" "la \ ua" "ua \ ub" "lb \ ub" using that lower_le_upper[of A] lower_le_upper[of B] by (auto simp: la_def ua_def lb_def ub_def set_of_eq) show ?thesis using ineqs by (simp add: set_of_subset_iff float_power_bnds_def max_def power_down_fl.rep_eq power_up_fl.rep_eq la_def[symmetric] ua_def[symmetric] lb_def[symmetric] ub_def[symmetric]) (auto intro!: power_down_mono power_up_mono intro: order_trans[where y=0]) qed lemma bounds_of_interval_eq_lower_upper: "bounds_of_interval ivl = (lower ivl, upper ivl)" if "lower ivl \ upper ivl" using that by (auto simp: lower.rep_eq upper.rep_eq) lemma real_interval_Ivl: "real_interval (Ivl a b) = Ivl a b" by transfer (auto simp: min_def) lemma set_of_mul_contains_real_zero: "0 \\<^sub>r (A * B)" if "0 \\<^sub>r A \ 0 \\<^sub>r B" using that set_of_mul_contains_zero[of A B] by (auto simp: set_of_eq) fun subdivide_interval :: "nat \ float interval \ float interval list" where "subdivide_interval 0 I = [I]" | "subdivide_interval (Suc n) I = ( let m = mid I in (subdivide_interval n (Ivl (lower I) m)) @ (subdivide_interval n (Ivl m (upper I))) )" lemma subdivide_interval_length: shows "length (subdivide_interval n I) = 2^n" by(induction n arbitrary: I, simp_all add: Let_def) lemma lower_le_mid: "lower x \ mid x" "real_of_float (lower x) \ mid x" and mid_le_upper: "mid x \ upper x" "real_of_float (mid x) \ upper x" unfolding mid_def subgoal by transfer auto subgoal by transfer auto subgoal by transfer auto subgoal by transfer auto done lemma subdivide_interval_correct: "list_ex (\i. x \\<^sub>r i) (subdivide_interval n I)" if "x \\<^sub>r I" for x::real using that proof(induction n arbitrary: x I) case 0 then show ?case by simp next case (Suc n) from \x \\<^sub>r I\ consider "x \\<^sub>r Ivl (lower I) (mid I)" | "x \\<^sub>r Ivl (mid I) (upper I)" by (cases "x \ real_of_float (mid I)") (auto simp: set_of_eq min_def lower_le_mid mid_le_upper) from this[case_names lower upper] show ?case by cases (use Suc.IH in \auto simp: Let_def\) qed fun interval_list_union :: "'a::lattice interval list \ 'a interval" where "interval_list_union [] = undefined" | "interval_list_union [I] = I" | "interval_list_union (I#Is) = sup I (interval_list_union Is)" lemma interval_list_union_correct: assumes "S \ []" assumes "i < length S" shows "set_of (S!i) \ set_of (interval_list_union S)" using assms proof(induction S arbitrary: i) case (Cons a S i) thus ?case proof(cases S) fix b S' assume "S = b # S'" hence "S \ []" by simp show ?thesis proof(cases i) case 0 show ?thesis apply(cases S) using interval_union_mono1 by (auto simp add: 0) next case (Suc i_prev) hence "i_prev < length S" using Cons(3) by simp from Cons(1)[OF \S \ []\ this] Cons(1) have "set_of ((a # S) ! i) \ set_of (interval_list_union S)" by (simp add: \i = Suc i_prev\) also have "... \ set_of (interval_list_union (a # S))" using \S \ []\ apply(cases S) using interval_union_mono2 by auto finally show ?thesis . qed qed simp qed simp lemma split_domain_correct: fixes x :: "real list" assumes "x all_in I" assumes split_correct: "\x a I. x \\<^sub>r I \ list_ex (\i::float interval. x \\<^sub>r i) (split I)" shows "list_ex (\s. x all_in s) (split_domain split I)" using assms(1) proof(induction I arbitrary: x) case (Cons I Is x) have "x \ []" using Cons(2) by auto obtain x' xs where x_decomp: "x = x' # xs" using \x \ []\ list.exhaust by auto hence "x' \\<^sub>r I" "xs all_in Is" using Cons(2) by auto show ?case using Cons(1)[OF \xs all_in Is\] split_correct[OF \x' \\<^sub>r I\] apply (auto simp add: list_ex_iff set_of_eq) by (smt length_Cons less_Suc_eq_0_disj nth_Cons_0 nth_Cons_Suc x_decomp) qed simp end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/Deterministic.thy b/thys/Transition_Systems_and_Automata/Automata/Deterministic.thy --- a/thys/Transition_Systems_and_Automata/Automata/Deterministic.thy +++ b/thys/Transition_Systems_and_Automata/Automata/Deterministic.thy @@ -1,559 +1,559 @@ theory Deterministic imports "../Transition_Systems/Transition_System" "../Transition_Systems/Transition_System_Extra" "../Transition_Systems/Transition_System_Construction" "../Basic/Degeneralization" begin type_synonym ('label, 'state) trans = "'label \ 'state \ 'state" (* TODO: is there a way to be less verbose in these locales? do we need to specify all the types? *) locale automaton = fixes automaton :: "'label set \ 'state \ ('label, 'state) trans \ 'condition \ 'automaton" fixes alphabet :: "'automaton \ 'label set" fixes initial :: "'automaton \ 'state" fixes transition :: "'automaton \ ('label, 'state) trans" fixes condition :: "'automaton \ 'condition" assumes automaton[simp]: "automaton (alphabet A) (initial A) (transition A) (condition A) = A" assumes alphabet[simp]: "alphabet (automaton a i t c) = a" assumes initial[simp]: "initial (automaton a i t c) = i" assumes transition[simp]: "transition (automaton a i t c) = t" assumes condition[simp]: "condition (automaton a i t c) = c" begin (* TODO: is there a way to use defines without renaming the constants? *) sublocale transition_system_initial "transition A" "\ a p. a \ alphabet A" "\ p. p = initial A" for A defines path' = path and run' = run and reachable' = reachable and nodes' = nodes by this lemma path_alt_def: "path A w p \ set w \ alphabet A" unfolding lists_iff_set[symmetric] proof show "w \ lists (alphabet A)" if "path A w p" using that by (induct arbitrary: p) (auto) show "path A w p" if "w \ lists (alphabet A)" using that by (induct arbitrary: p) (auto) qed lemma run_alt_def: "run A w p \ sset w \ alphabet A" unfolding streams_iff_sset[symmetric] proof show "w \ streams (alphabet A)" if "run A w p" using that by (coinduction arbitrary: w p) (force elim: run.cases) show "run A w p" if "w \ streams (alphabet A)" using that by (coinduction arbitrary: w p) (force elim: streams.cases) qed end (* TODO: create analogous locale for DFAs (automaton_target) *) locale automaton_trace = automaton automaton alphabet initial transition condition for automaton :: "'label set \ 'state \ ('label, 'state) trans \ 'condition \ 'automaton" and alphabet :: "'automaton \ 'label set" and initial :: "'automaton \ 'state" and transition :: "'automaton \ ('label, 'state) trans" and condition :: "'automaton \ 'condition" + fixes test :: "'condition \ 'state stream \ bool" begin definition language :: "'automaton \ 'label stream set" where "language A \ {w. run A w (initial A) \ test (condition A) (trace A w (initial A))}" lemma language[intro]: assumes "run A w (initial A)" "test (condition A) (trace A w (initial A))" shows "w \ language A" using assms unfolding language_def by auto lemma language_elim[elim]: assumes "w \ language A" obtains "run A w (initial A)" "test (condition A) (trace A w (initial A))" using assms unfolding language_def by auto lemma language_alphabet: "language A \ streams (alphabet A)" unfolding language_def run_alt_def using sset_streams by auto end locale automaton_degeneralization = a: automaton automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 + b: automaton automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 for automaton\<^sub>1 :: "'label set \ 'state \ ('label, 'state) trans \ 'state pred gen \ 'automaton\<^sub>1" and alphabet\<^sub>1 :: "'automaton\<^sub>1 \ 'label set" and initial\<^sub>1 :: "'automaton\<^sub>1 \ 'state" and transition\<^sub>1 :: "'automaton\<^sub>1 \ ('label, 'state) trans" and condition\<^sub>1 :: "'automaton\<^sub>1 \ 'state pred gen" and automaton\<^sub>2 :: "'label set \ 'state degen \ ('label, 'state degen) trans \ 'state degen pred \ 'automaton\<^sub>2" and alphabet\<^sub>2 :: "'automaton\<^sub>2 \ 'label set" and initial\<^sub>2 :: "'automaton\<^sub>2 \ 'state degen" and transition\<^sub>2 :: "'automaton\<^sub>2 \ ('label, 'state degen) trans" and condition\<^sub>2 :: "'automaton\<^sub>2 \ 'state degen pred" begin definition degeneralize :: "'automaton\<^sub>1 \ 'automaton\<^sub>2" where "degeneralize A \ automaton\<^sub>2 (alphabet\<^sub>1 A) (initial\<^sub>1 A, 0) (\ a (p, k). (transition\<^sub>1 A a p, count (condition\<^sub>1 A) p k)) (degen (condition\<^sub>1 A))" lemma degeneralize_simps[simp]: "alphabet\<^sub>2 (degeneralize A) = alphabet\<^sub>1 A" "initial\<^sub>2 (degeneralize A) = (initial\<^sub>1 A, 0)" "transition\<^sub>2 (degeneralize A) a (p, k) = (transition\<^sub>1 A a p, count (condition\<^sub>1 A) p k)" "condition\<^sub>2 (degeneralize A) = degen (condition\<^sub>1 A)" unfolding degeneralize_def by auto lemma degeneralize_target[simp]: "b.target (degeneralize A) w (p, k) = (a.target A w p, fold (count (condition\<^sub>1 A)) (butlast (p # a.states A w p)) k)" by (induct w arbitrary: p k) (auto) lemma degeneralize_states[simp]: "b.states (degeneralize A) w (p, k) = a.states A w p || scan (count (condition\<^sub>1 A)) (p # a.states A w p) k" by (induct w arbitrary: p k) (auto) lemma degeneralize_trace[simp]: "b.trace (degeneralize A) w (p, k) = a.trace A w p ||| sscan (count (condition\<^sub>1 A)) (p ## a.trace A w p) k" by (coinduction arbitrary: w p k) (auto) lemma degeneralize_path[iff]: "b.path (degeneralize A) w (p, k) \ a.path A w p" unfolding a.path_alt_def b.path_alt_def by simp lemma degeneralize_run[iff]: "b.run (degeneralize A) w (p, k) \ a.run A w p" unfolding a.run_alt_def b.run_alt_def by simp lemma degeneralize_reachable_fst[simp]: "fst ` b.reachable (degeneralize A) (p, k) = a.reachable A p" unfolding a.reachable_alt_def b.reachable_alt_def image_def by simp lemma degeneralize_reachable_snd_empty[simp]: assumes "condition\<^sub>1 A = []" shows "snd ` b.reachable (degeneralize A) (p, k) = {k}" proof - have "snd ql = k" if "ql \ b.reachable (degeneralize A) (p, k)" for ql using that assms by induct auto then show ?thesis by auto qed lemma degeneralize_reachable_empty[simp]: assumes "condition\<^sub>1 A = []" shows "b.reachable (degeneralize A) (p, k) = a.reachable A p \ {k}" using degeneralize_reachable_fst degeneralize_reachable_snd_empty assms by (metis prod_singleton(2)) lemma degeneralize_reachable_snd: "snd ` b.reachable (degeneralize A) (p, k) \ insert k {0 ..< length (condition\<^sub>1 A)}" by (cases "condition\<^sub>1 A = []") (auto) lemma degeneralize_reachable: "b.reachable (degeneralize A) (p, k) \ a.reachable A p \ insert k {0 ..< length (condition\<^sub>1 A)}" by (cases "condition\<^sub>1 A = []") (auto 0 3) lemma degeneralize_nodes_fst[simp]: "fst ` b.nodes (degeneralize A) = a.nodes A" unfolding a.nodes_alt_def b.nodes_alt_def by simp lemma degeneralize_nodes_snd_empty: assumes "condition\<^sub>1 A = []" shows "snd ` b.nodes (degeneralize A) = {0}" using assms unfolding b.nodes_alt_def by auto lemma degeneralize_nodes_empty: assumes "condition\<^sub>1 A = []" shows "b.nodes (degeneralize A) = a.nodes A \ {0}" using assms unfolding b.nodes_alt_def by auto lemma degeneralize_nodes_snd: "snd ` b.nodes (degeneralize A) \ insert 0 {0 ..< length (condition\<^sub>1 A)}" using degeneralize_reachable_snd unfolding b.nodes_alt_def by auto lemma degeneralize_nodes: "b.nodes (degeneralize A) \ a.nodes A \ insert 0 {0 ..< length (condition\<^sub>1 A)}" using degeneralize_reachable unfolding a.nodes_alt_def b.nodes_alt_def by simp lemma degeneralize_nodes_finite[iff]: "finite (b.nodes (degeneralize A)) \ finite (a.nodes A)" proof show "finite (a.nodes A)" if "finite (b.nodes (degeneralize A))" using that by (auto simp flip: degeneralize_nodes_fst) show "finite (b.nodes (degeneralize A))" if "finite (a.nodes A)" - using degeneralize_nodes that finite_subset by fastforce + using finite_subset degeneralize_nodes that by blast qed lemma degeneralize_nodes_card: "card (b.nodes (degeneralize A)) \ max 1 (length (condition\<^sub>1 A)) * card (a.nodes A)" proof (cases "finite (a.nodes A)") case True have "card (b.nodes (degeneralize A)) \ card (a.nodes A \ insert 0 {0 ..< length (condition\<^sub>1 A)})" using degeneralize_nodes True by (blast intro: card_mono) also have "\ = card (insert 0 {0 ..< length (condition\<^sub>1 A)}) * card (a.nodes A)" unfolding card_cartesian_product by simp also have "card (insert 0 {0 ..< length (condition\<^sub>1 A)}) = max 1 (length (condition\<^sub>1 A))" by (simp add: card_insert_if Suc_leI max_absorb2) finally show ?thesis by this next case False then have "card (a.nodes A) = 0" "card (b.nodes (degeneralize A)) = 0" by auto then show ?thesis by simp qed end locale automaton_degeneralization_trace = automaton_degeneralization automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 + a: automaton_trace automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_trace automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 for automaton\<^sub>1 :: "'label set \ 'state \ ('label, 'state) trans \ 'state pred gen \ 'automaton\<^sub>1" and alphabet\<^sub>1 :: "'automaton\<^sub>1 \ 'label set" and initial\<^sub>1 :: "'automaton\<^sub>1 \ 'state" and transition\<^sub>1 :: "'automaton\<^sub>1 \ ('label, 'state) trans" and condition\<^sub>1 :: "'automaton\<^sub>1 \ 'state pred gen" and test\<^sub>1 :: "'state pred gen \ 'state stream \ bool" and automaton\<^sub>2 :: "'label set \ 'state degen \ ('label, 'state degen) trans \ 'state degen pred \ 'automaton\<^sub>2" and alphabet\<^sub>2 :: "'automaton\<^sub>2 \ 'label set" and initial\<^sub>2 :: "'automaton\<^sub>2 \ 'state degen" and transition\<^sub>2 :: "'automaton\<^sub>2 \ ('label, 'state degen) trans" and condition\<^sub>2 :: "'automaton\<^sub>2 \ 'state degen pred" and test\<^sub>2 :: "'state degen pred \ 'state degen stream \ bool" + assumes test[iff]: "test\<^sub>2 (degen cs) (r ||| sscan (count cs) (p ## r) k) \ test\<^sub>1 cs r" begin lemma degeneralize_language[simp]: "b.language (degeneralize A) = a.language A" by (force simp del: sscan_scons) end locale automaton_combination = a: automaton automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 + b: automaton automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 + c: automaton automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 for automaton\<^sub>1 :: "'label set \ 'state\<^sub>1 \ ('label, 'state\<^sub>1) trans \ 'condition\<^sub>1 \ 'automaton\<^sub>1" and alphabet\<^sub>1 :: "'automaton\<^sub>1 \ 'label set" and initial\<^sub>1 :: "'automaton\<^sub>1 \ 'state\<^sub>1" and transition\<^sub>1 :: "'automaton\<^sub>1 \ ('label, 'state\<^sub>1) trans" and condition\<^sub>1 :: "'automaton\<^sub>1 \ 'condition\<^sub>1" and automaton\<^sub>2 :: "'label set \ 'state\<^sub>2 \ ('label, 'state\<^sub>2) trans \ 'condition\<^sub>2 \ 'automaton\<^sub>2" and alphabet\<^sub>2 :: "'automaton\<^sub>2 \ 'label set" and initial\<^sub>2 :: "'automaton\<^sub>2 \ 'state\<^sub>2" and transition\<^sub>2 :: "'automaton\<^sub>2 \ ('label, 'state\<^sub>2) trans" and condition\<^sub>2 :: "'automaton\<^sub>2 \ 'condition\<^sub>2" and automaton\<^sub>3 :: "'label set \ 'state\<^sub>1 \ 'state\<^sub>2 \ ('label, 'state\<^sub>1 \ 'state\<^sub>2) trans \ 'condition\<^sub>3 \ 'automaton\<^sub>3" and alphabet\<^sub>3 :: "'automaton\<^sub>3 \ 'label set" and initial\<^sub>3 :: "'automaton\<^sub>3 \ 'state\<^sub>1 \ 'state\<^sub>2" and transition\<^sub>3 :: "'automaton\<^sub>3 \ ('label, 'state\<^sub>1 \ 'state\<^sub>2) trans" and condition\<^sub>3 :: "'automaton\<^sub>3 \ 'condition\<^sub>3" + fixes condition :: "'condition\<^sub>1 \ 'condition\<^sub>2 \ 'condition\<^sub>3" begin definition combine :: "'automaton\<^sub>1 \ 'automaton\<^sub>2 \ 'automaton\<^sub>3" where "combine A B \ automaton\<^sub>3 (alphabet\<^sub>1 A \ alphabet\<^sub>2 B) (initial\<^sub>1 A, initial\<^sub>2 B) (\ a (p, q). (transition\<^sub>1 A a p, transition\<^sub>2 B a q)) (condition (condition\<^sub>1 A) (condition\<^sub>2 B))" lemma combine_simps[simp]: "alphabet\<^sub>3 (combine A B) = alphabet\<^sub>1 A \ alphabet\<^sub>2 B" "initial\<^sub>3 (combine A B) = (initial\<^sub>1 A, initial\<^sub>2 B)" "transition\<^sub>3 (combine A B) a (p, q) = (transition\<^sub>1 A a p, transition\<^sub>2 B a q)" "condition\<^sub>3 (combine A B) = condition (condition\<^sub>1 A) (condition\<^sub>2 B)" unfolding combine_def by auto lemma combine_target[simp]: "c.target (combine A B) w (p, q) = (a.target A w p, b.target B w q)" by (induct w arbitrary: p q) (auto) lemma combine_states[simp]: "c.states (combine A B) w (p, q) = a.states A w p || b.states B w q" by (induct w arbitrary: p q) (auto) lemma combine_trace[simp]: "c.trace (combine A B) w (p, q) = a.trace A w p ||| b.trace B w q" by (coinduction arbitrary: w p q) (auto) lemma combine_path[iff]: "c.path (combine A B) w (p, q) \ a.path A w p \ b.path B w q" unfolding a.path_alt_def b.path_alt_def c.path_alt_def by simp lemma combine_run[iff]: "c.run (combine A B) w (p, q) \ a.run A w p \ b.run B w q" unfolding a.run_alt_def b.run_alt_def c.run_alt_def by simp lemma combine_reachable[simp]: "c.reachable (combine A B) (p, q) \ a.reachable A p \ b.reachable B q" unfolding c.reachable_alt_def by auto lemma combine_nodes[simp]: "c.nodes (combine A B) \ a.nodes A \ b.nodes B" unfolding a.nodes_alt_def b.nodes_alt_def c.nodes_alt_def by auto lemma combine_reachable_fst[simp]: assumes "alphabet\<^sub>1 A \ alphabet\<^sub>2 B" shows "fst ` c.reachable (combine A B) (p, q) = a.reachable A p" using assms unfolding a.reachable_alt_def a.path_alt_def unfolding b.reachable_alt_def b.path_alt_def unfolding c.reachable_alt_def c.path_alt_def by auto force lemma combine_reachable_snd[simp]: assumes "alphabet\<^sub>1 A \ alphabet\<^sub>2 B" shows "snd ` c.reachable (combine A B) (p, q) = b.reachable B q" using assms unfolding a.reachable_alt_def a.path_alt_def unfolding b.reachable_alt_def b.path_alt_def unfolding c.reachable_alt_def c.path_alt_def by auto force lemma combine_nodes_fst[simp]: assumes "alphabet\<^sub>1 A \ alphabet\<^sub>2 B" shows "fst ` c.nodes (combine A B) = a.nodes A" using assms combine_reachable_fst unfolding a.nodes_alt_def b.nodes_alt_def c.nodes_alt_def by fastforce lemma combine_nodes_snd[simp]: assumes "alphabet\<^sub>1 A \ alphabet\<^sub>2 B" shows "snd ` c.nodes (combine A B) = b.nodes B" using assms combine_reachable_snd unfolding a.nodes_alt_def b.nodes_alt_def c.nodes_alt_def by fastforce lemma combine_nodes_finite[intro]: assumes "finite (a.nodes A)" "finite (b.nodes B)" shows "finite (c.nodes (combine A B))" proof (rule finite_subset) show "c.nodes (combine A B) \ a.nodes A \ b.nodes B" using combine_nodes by this show "finite (a.nodes A \ b.nodes B)" using assms by simp qed lemma combine_nodes_finite_strong[iff]: assumes "alphabet\<^sub>1 A = alphabet\<^sub>2 B" shows "finite (c.nodes (combine A B)) \ finite (a.nodes A) \ finite (b.nodes B)" proof safe show "finite (a.nodes A)" if "finite (c.nodes (combine A B))" using combine_nodes_fst assms that by (metis finite_imageI equalityD1) show "finite (b.nodes B)" if "finite (c.nodes (combine A B))" using combine_nodes_snd assms that by (metis finite_imageI equalityD2) show "finite (c.nodes (combine A B))" if "finite (a.nodes A)" "finite (b.nodes B)" using that by rule qed lemma combine_nodes_card[intro]: assumes "finite (a.nodes A)" "finite (b.nodes B)" shows "card (c.nodes (combine A B)) \ card (a.nodes A) * card (b.nodes B)" proof - have "card (c.nodes (combine A B)) \ card (a.nodes A \ b.nodes B)" proof (rule card_mono) show "finite (a.nodes A \ b.nodes B)" using assms by simp show "c.nodes (combine A B) \ a.nodes A \ b.nodes B" using combine_nodes by this qed also have "\ = card (a.nodes A) * card (b.nodes B)" using card_cartesian_product by this finally show ?thesis by this qed lemma combine_nodes_card_strong[intro]: assumes "alphabet\<^sub>1 A = alphabet\<^sub>2 B" shows "card (c.nodes (combine A B)) \ card (a.nodes A) * card (b.nodes B)" proof (cases "finite (a.nodes A) \ finite (b.nodes B)") case True show ?thesis using True by auto next case False have 1: "card (c.nodes (combine A B)) = 0" using False assms by simp have 2: "card (a.nodes A) * card (b.nodes B) = 0" using False by auto show ?thesis using 1 2 by simp qed end locale automaton_intersection_trace = automaton_combination automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 condition + a: automaton_trace automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_trace automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 + c: automaton_trace automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 test\<^sub>3 for automaton\<^sub>1 :: "'label set \ 'state\<^sub>1 \ ('label, 'state\<^sub>1) trans \ 'condition\<^sub>1 \ 'automaton\<^sub>1" and alphabet\<^sub>1 :: "'automaton\<^sub>1 \ 'label set" and initial\<^sub>1 :: "'automaton\<^sub>1 \ 'state\<^sub>1" and transition\<^sub>1 :: "'automaton\<^sub>1 \ ('label, 'state\<^sub>1) trans" and condition\<^sub>1 :: "'automaton\<^sub>1 \ 'condition\<^sub>1" and test\<^sub>1 :: "'condition\<^sub>1 \ 'state\<^sub>1 stream \ bool" and automaton\<^sub>2 :: "'label set \ 'state\<^sub>2 \ ('label, 'state\<^sub>2) trans \ 'condition\<^sub>2 \ 'automaton\<^sub>2" and alphabet\<^sub>2 :: "'automaton\<^sub>2 \ 'label set" and initial\<^sub>2 :: "'automaton\<^sub>2 \ 'state\<^sub>2" and transition\<^sub>2 :: "'automaton\<^sub>2 \ ('label, 'state\<^sub>2) trans" and condition\<^sub>2 :: "'automaton\<^sub>2 \ 'condition\<^sub>2" and test\<^sub>2 :: "'condition\<^sub>2 \ 'state\<^sub>2 stream \ bool" and automaton\<^sub>3 :: "'label set \ 'state\<^sub>1 \ 'state\<^sub>2 \ ('label, 'state\<^sub>1 \ 'state\<^sub>2) trans \ 'condition\<^sub>3 \ 'automaton\<^sub>3" and alphabet\<^sub>3 :: "'automaton\<^sub>3 \ 'label set" and initial\<^sub>3 :: "'automaton\<^sub>3 \ 'state\<^sub>1 \ 'state\<^sub>2" and transition\<^sub>3 :: "'automaton\<^sub>3 \ ('label, 'state\<^sub>1 \ 'state\<^sub>2) trans" and condition\<^sub>3 :: "'automaton\<^sub>3 \ 'condition\<^sub>3" and test\<^sub>3 :: "'condition\<^sub>3 \ ('state\<^sub>1 \ 'state\<^sub>2) stream \ bool" and condition :: "'condition\<^sub>1 \ 'condition\<^sub>2 \ 'condition\<^sub>3" + assumes test[iff]: "test\<^sub>3 (condition c\<^sub>1 c\<^sub>2) (u ||| v) \ test\<^sub>1 c\<^sub>1 u \ test\<^sub>2 c\<^sub>2 v" begin lemma combine_language[simp]: "c.language (combine A B) = a.language A \ b.language B" by force end locale automaton_union_trace = automaton_combination automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 condition + a: automaton_trace automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_trace automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 + c: automaton_trace automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 test\<^sub>3 for automaton\<^sub>1 :: "'label set \ 'state\<^sub>1 \ ('label, 'state\<^sub>1) trans \ 'condition\<^sub>1 \ 'automaton\<^sub>1" and alphabet\<^sub>1 :: "'automaton\<^sub>1 \ 'label set" and initial\<^sub>1 :: "'automaton\<^sub>1 \ 'state\<^sub>1" and transition\<^sub>1 :: "'automaton\<^sub>1 \ ('label, 'state\<^sub>1) trans" and condition\<^sub>1 :: "'automaton\<^sub>1 \ 'condition\<^sub>1" and test\<^sub>1 :: "'condition\<^sub>1 \ 'state\<^sub>1 stream \ bool" and automaton\<^sub>2 :: "'label set \ 'state\<^sub>2 \ ('label, 'state\<^sub>2) trans \ 'condition\<^sub>2 \ 'automaton\<^sub>2" and alphabet\<^sub>2 :: "'automaton\<^sub>2 \ 'label set" and initial\<^sub>2 :: "'automaton\<^sub>2 \ 'state\<^sub>2" and transition\<^sub>2 :: "'automaton\<^sub>2 \ ('label, 'state\<^sub>2) trans" and condition\<^sub>2 :: "'automaton\<^sub>2 \ 'condition\<^sub>2" and test\<^sub>2 :: "'condition\<^sub>2 \ 'state\<^sub>2 stream \ bool" and automaton\<^sub>3 :: "'label set \ 'state\<^sub>1 \ 'state\<^sub>2 \ ('label, 'state\<^sub>1 \ 'state\<^sub>2) trans \ 'condition\<^sub>3 \ 'automaton\<^sub>3" and alphabet\<^sub>3 :: "'automaton\<^sub>3 \ 'label set" and initial\<^sub>3 :: "'automaton\<^sub>3 \ 'state\<^sub>1 \ 'state\<^sub>2" and transition\<^sub>3 :: "'automaton\<^sub>3 \ ('label, 'state\<^sub>1 \ 'state\<^sub>2) trans" and condition\<^sub>3 :: "'automaton\<^sub>3 \ 'condition\<^sub>3" and test\<^sub>3 :: "'condition\<^sub>3 \ ('state\<^sub>1 \ 'state\<^sub>2) stream \ bool" and condition :: "'condition\<^sub>1 \ 'condition\<^sub>2 \ 'condition\<^sub>3" + assumes test[iff]: "test\<^sub>3 (condition c\<^sub>1 c\<^sub>2) (u ||| v) \ test\<^sub>1 c\<^sub>1 u \ test\<^sub>2 c\<^sub>2 v" begin lemma combine_language[simp]: assumes "alphabet\<^sub>1 A = alphabet\<^sub>2 B" shows "c.language (combine A B) = a.language A \ b.language B" using assms by (force simp: a.run_alt_def b.run_alt_def) end (* TODO: complete this in analogy to the pair case *) locale automaton_combination_list = a: automaton automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 + b: automaton automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 for automaton\<^sub>1 :: "'label set \ 'state \ ('label, 'state) trans \ 'condition\<^sub>1 \ 'automaton\<^sub>1" and alphabet\<^sub>1 :: "'automaton\<^sub>1 \ 'label set" and initial\<^sub>1 :: "'automaton\<^sub>1 \ 'state" and transition\<^sub>1 :: "'automaton\<^sub>1 \ ('label, 'state) trans" and condition\<^sub>1 :: "'automaton\<^sub>1 \ 'condition\<^sub>1" and automaton\<^sub>2 :: "'label set \ 'state list \ ('label, 'state list) trans \ 'condition\<^sub>2 \ 'automaton\<^sub>2" and alphabet\<^sub>2 :: "'automaton\<^sub>2 \ 'label set" and initial\<^sub>2 :: "'automaton\<^sub>2 \ 'state list" and transition\<^sub>2 :: "'automaton\<^sub>2 \ ('label, 'state list) trans" and condition\<^sub>2 :: "'automaton\<^sub>2 \ 'condition\<^sub>2" + fixes condition :: "'condition\<^sub>1 list \ 'condition\<^sub>2" begin definition combine :: "'automaton\<^sub>1 list \ 'automaton\<^sub>2" where "combine AA \ automaton\<^sub>2 (\ (alphabet\<^sub>1 ` set AA)) (map initial\<^sub>1 AA) (\ a pp. map2 (\ A p. transition\<^sub>1 A a p) AA pp) (condition (map condition\<^sub>1 AA))" lemma combine_simps[simp]: "alphabet\<^sub>2 (combine AA) = \ (alphabet\<^sub>1 ` set AA)" "initial\<^sub>2 (combine AA) = map initial\<^sub>1 AA" "transition\<^sub>2 (combine AA) a pp = map2 (\ A p. transition\<^sub>1 A a p) AA pp" "condition\<^sub>2 (combine AA) = condition (map condition\<^sub>1 AA)" unfolding combine_def by auto (* TODO: get rid of indices, express this using stranspose *) lemma combine_trace_smap: assumes "length pp = length AA" "k < length AA" shows "smap (\ pp. pp ! k) (b.trace (combine AA) w pp) = a.trace (AA ! k) w (pp ! k)" using assms by (coinduction arbitrary: w pp) (force) lemma combine_nodes_length: assumes "pp \ b.nodes (combine AA)" shows "length pp = length AA" using assms by induct auto lemma combine_nodes[intro]: assumes "pp \ b.nodes (combine AA)" "k < length pp" shows "pp ! k \ a.nodes (AA ! k)" using assms by induct auto lemma combine_nodes_finite[intro]: assumes "list_all (finite \ a.nodes) AA" shows "finite (b.nodes (combine AA))" proof (rule finite_subset) show "b.nodes (combine AA) \ listset (map a.nodes AA)" by (force simp: listset_member list_all2_conv_all_nth combine_nodes_length) have "finite (listset (map a.nodes AA)) \ list_all finite (map a.nodes AA)" by (rule listset_finite) (auto simp: list_all_iff) then show "finite (listset (map a.nodes AA))" using assms by (simp add: list.pred_map) qed lemma combine_nodes_card: assumes "list_all (finite \ a.nodes) AA" shows "card (b.nodes (combine AA)) \ prod_list (map (card \ a.nodes) AA)" proof - have "card (b.nodes (combine AA)) \ card (listset (map a.nodes AA))" proof (rule card_mono) have "finite (listset (map a.nodes AA)) \ list_all finite (map a.nodes AA)" by (rule listset_finite) (auto simp: list_all_iff) then show "finite (listset (map a.nodes AA))" using assms by (simp add: list.pred_map) show "b.nodes (combine AA) \ listset (map a.nodes AA)" by (force simp: listset_member list_all2_conv_all_nth combine_nodes_length) qed also have "\ = prod_list (map (card \ a.nodes) AA)" by simp finally show ?thesis by this qed end locale automaton_intersection_list_trace = automaton_combination_list automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 condition + a: automaton_trace automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_trace automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 for automaton\<^sub>1 :: "'label set \ 'state \ ('label, 'state) trans \ 'condition\<^sub>1 \ 'automaton\<^sub>1" and alphabet\<^sub>1 :: "'automaton\<^sub>1 \ 'label set" and initial\<^sub>1 :: "'automaton\<^sub>1 \ 'state" and transition\<^sub>1 :: "'automaton\<^sub>1 \ ('label, 'state) trans" and condition\<^sub>1 :: "'automaton\<^sub>1 \ 'condition\<^sub>1" and test\<^sub>1 :: "'condition\<^sub>1 \ 'state stream \ bool" and automaton\<^sub>2 :: "'label set \ 'state list \ ('label, 'state list) trans \ 'condition\<^sub>2 \ 'automaton\<^sub>2" and alphabet\<^sub>2 :: "'automaton\<^sub>2 \ 'label set" and initial\<^sub>2 :: "'automaton\<^sub>2 \ 'state list" and transition\<^sub>2 :: "'automaton\<^sub>2 \ ('label, 'state list) trans" and condition\<^sub>2 :: "'automaton\<^sub>2 \ 'condition\<^sub>2" and test\<^sub>2 :: "'condition\<^sub>2 \ 'state list stream \ bool" and condition :: "'condition\<^sub>1 list \ 'condition\<^sub>2" + assumes test[iff]: "test\<^sub>2 (condition cs) w \ (\ k < length cs. test\<^sub>1 (cs ! k) (smap (\ pp. pp ! k) w))" begin lemma combine_language[simp]: "b.language (combine AA) = \ (a.language ` set AA)" unfolding a.language_def b.language_def unfolding a.run_alt_def b.run_alt_def by (auto simp: combine_trace_smap iff: in_set_conv_nth) end locale automaton_union_list_trace = automaton_combination_list automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 condition + a: automaton_trace automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_trace automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 for automaton\<^sub>1 :: "'label set \ 'state \ ('label, 'state) trans \ 'condition\<^sub>1 \ 'automaton\<^sub>1" and alphabet\<^sub>1 :: "'automaton\<^sub>1 \ 'label set" and initial\<^sub>1 :: "'automaton\<^sub>1 \ 'state" and transition\<^sub>1 :: "'automaton\<^sub>1 \ ('label, 'state) trans" and condition\<^sub>1 :: "'automaton\<^sub>1 \ 'condition\<^sub>1" and test\<^sub>1 :: "'condition\<^sub>1 \ 'state stream \ bool" and automaton\<^sub>2 :: "'label set \ 'state list \ ('label, 'state list) trans \ 'condition\<^sub>2 \ 'automaton\<^sub>2" and alphabet\<^sub>2 :: "'automaton\<^sub>2 \ 'label set" and initial\<^sub>2 :: "'automaton\<^sub>2 \ 'state list" and transition\<^sub>2 :: "'automaton\<^sub>2 \ ('label, 'state list) trans" and condition\<^sub>2 :: "'automaton\<^sub>2 \ 'condition\<^sub>2" and test\<^sub>2 :: "'condition\<^sub>2 \ 'state list stream \ bool" and condition :: "'condition\<^sub>1 list \ 'condition\<^sub>2" + assumes test[iff]: "test\<^sub>2 (condition cs) w \ (\ k < length cs. test\<^sub>1 (cs ! k) (smap (\ pp. pp ! k) w))" begin lemma combine_language[simp]: assumes "\ (alphabet\<^sub>1 ` set AA) = \ (alphabet\<^sub>1 ` set AA)" shows "b.language (combine AA) = \ (a.language ` set AA)" using assms unfolding a.language_def b.language_def unfolding a.run_alt_def b.run_alt_def by (auto simp: combine_trace_smap iff: in_set_conv_nth) (metis INT_subset_iff in_set_conv_nth) end end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Combine.thy b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Combine.thy --- a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Combine.thy +++ b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Combine.thy @@ -1,40 +1,46 @@ section \Nondeterministic Büchi Automata Combinations\ theory NBA_Combine imports NBA NGBA begin global_interpretation degeneralization: automaton_degeneralization_trace ngba ngba.alphabet ngba.initial ngba.transition ngba.accepting "gen infs" nba nba.alphabet nba.initial nba.transition nba.accepting infs defines degeneralize = degeneralization.degeneralize by unfold_locales auto lemmas degeneralize_language[simp] = degeneralization.degeneralize_language[folded NBA.language_def] + lemmas degeneralize_nodes_finite[iff] = degeneralization.degeneralize_nodes_finite[folded NBA.nodes_def] global_interpretation intersection: automaton_intersection_trace nba nba.alphabet nba.initial nba.transition nba.accepting infs nba nba.alphabet nba.initial nba.transition nba.accepting infs ngba ngba.alphabet ngba.initial ngba.transition ngba.accepting "gen infs" "\ c\<^sub>1 c\<^sub>2. [c\<^sub>1 \ fst, c\<^sub>2 \ snd]" defines intersect' = intersection.intersect by unfold_locales auto - lemmas intersect'_language[simp] = intersection.combine_language[folded NGBA.language_def] + lemmas intersect'_language[simp] = intersection.intersect_language[folded NGBA.language_def] + lemmas intersect'_nodes_finite[intro] = intersection.intersect_nodes_finite[folded NGBA.nodes_def] global_interpretation union: automaton_union_trace nba nba.alphabet nba.initial nba.transition nba.accepting infs nba nba.alphabet nba.initial nba.transition nba.accepting infs nba nba.alphabet nba.initial nba.transition nba.accepting infs case_sum defines union = union.union by (unfold_locales) (auto simp: comp_def) lemmas union_language = union.union_language abbreviation intersect where "intersect A B \ degeneralize (intersect' A B)" lemma intersect_language[simp]: "NBA.language (intersect A B) = NBA.language A \ NBA.language B" by simp + lemma intersect_nodes_finite[intro]: + assumes "finite (NBA.nodes A)" "finite (NBA.nodes B)" + shows "finite (NBA.nodes (intersect A B))" + using intersect'_nodes_finite assms by simp end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Explicit.thy b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Explicit.thy --- a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Explicit.thy +++ b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Explicit.thy @@ -1,162 +1,170 @@ section \Explicit Nondeterministic Büchi Automata\ theory NBA_Explicit imports NBA_Algorithms begin datatype ('label, 'state) nbae = nbae (alphabete: "'label set") (initiale: "'state set") (transitione: "('state \ 'label \ 'state) set") (acceptinge: "'state set") definition nbae_rel where [to_relAPP]: "nbae_rel L S \ {(A\<^sub>1, A\<^sub>2). (alphabete A\<^sub>1, alphabete A\<^sub>2) \ \L\ set_rel \ (initiale A\<^sub>1, initiale A\<^sub>2) \ \S\ set_rel \ (transitione A\<^sub>1, transitione A\<^sub>2) \ \S \\<^sub>r L \\<^sub>r S\ set_rel \ (acceptinge A\<^sub>1, acceptinge A\<^sub>2) \ \S\ set_rel}" lemma nbae_param[param, autoref_rules]: "(nbae, nbae) \ \L\ set_rel \ \S\ set_rel \ \S \\<^sub>r L \\<^sub>r S\ set_rel \ \S\ set_rel \ \L, S\ nbae_rel" "(alphabete, alphabete) \ \L, S\ nbae_rel \ \L\ set_rel" "(initiale, initiale) \ \L, S\ nbae_rel \ \S\ set_rel" "(transitione, transitione) \ \L, S\ nbae_rel \ \S \\<^sub>r L \\<^sub>r S\ set_rel" "(acceptinge, acceptinge) \ \L, S\ nbae_rel \ \S\ set_rel" unfolding nbae_rel_def by auto lemma nbae_rel_id[simp]: "\Id, Id\ nbae_rel = Id" unfolding nbae_rel_def using nbae.expand by auto lemma nbae_rel_comp[simp]: "\L\<^sub>1 O L\<^sub>2, S\<^sub>1 O S\<^sub>2\ nbae_rel = \L\<^sub>1, S\<^sub>1\ nbae_rel O \L\<^sub>2, S\<^sub>2\ nbae_rel" proof safe fix A B assume 1: "(A, B) \ \L\<^sub>1 O L\<^sub>2, S\<^sub>1 O S\<^sub>2\ nbae_rel" obtain a b c d where 2: "(alphabete A, a) \ \L\<^sub>1\ set_rel" "(a, alphabete B) \ \L\<^sub>2\ set_rel" "(initiale A, b) \ \S\<^sub>1\ set_rel" "(b, initiale B) \ \S\<^sub>2\ set_rel" "(transitione A, c) \ \S\<^sub>1 \\<^sub>r L\<^sub>1 \\<^sub>r S\<^sub>1\ set_rel" "(c, transitione B) \ \S\<^sub>2 \\<^sub>r L\<^sub>2 \\<^sub>r S\<^sub>2\ set_rel" "(acceptinge A, d) \ \S\<^sub>1\ set_rel" "(d, acceptinge B) \ \S\<^sub>2\ set_rel" using 1 unfolding nbae_rel_def prod_rel_compp set_rel_compp by auto show "(A, B) \ \L\<^sub>1, S\<^sub>1\ nbae_rel O \L\<^sub>2, S\<^sub>2\ nbae_rel" proof show "(A, nbae a b c d) \ \L\<^sub>1, S\<^sub>1\ nbae_rel" using 2 unfolding nbae_rel_def by auto show "(nbae a b c d, B) \ \L\<^sub>2, S\<^sub>2\ nbae_rel" using 2 unfolding nbae_rel_def by auto qed next show "(A, C) \ \L\<^sub>1 O L\<^sub>2, S\<^sub>1 O S\<^sub>2\ nbae_rel" if "(A, B) \ \L\<^sub>1, S\<^sub>1\ nbae_rel" "(B, C) \ \L\<^sub>2, S\<^sub>2\ nbae_rel" for A B C using that unfolding nbae_rel_def prod_rel_compp set_rel_compp by auto qed (* TODO: why do we need all this setup? can't i_of_rel do the trick? *) consts i_nbae_scheme :: "interface \ interface \ interface" context begin interpretation autoref_syn by this lemma nbae_scheme_itype[autoref_itype]: "nbae ::\<^sub>i \L\\<^sub>i i_set \\<^sub>i \S\\<^sub>i i_set \\<^sub>i \\S, \L, S\\<^sub>i i_prod\\<^sub>i i_prod\\<^sub>i i_set \\<^sub>i \S\\<^sub>i i_set \\<^sub>i \L, S\\<^sub>i i_nbae_scheme" "alphabete ::\<^sub>i \L, S\\<^sub>i i_nbae_scheme \\<^sub>i \L\\<^sub>i i_set" "initiale ::\<^sub>i \L, S\\<^sub>i i_nbae_scheme \\<^sub>i \S\\<^sub>i i_set" "transitione ::\<^sub>i \L, S\\<^sub>i i_nbae_scheme \\<^sub>i \\S, \L, S\\<^sub>i i_prod\\<^sub>i i_prod\\<^sub>i i_set" "acceptinge ::\<^sub>i \L, S\\<^sub>i i_nbae_scheme \\<^sub>i \S\\<^sub>i i_set" by auto end datatype ('label, 'state) nbaei = nbaei (alphabetei: "'label list") (initialei: "'state list") (transitionei: "('state \ 'label \ 'state) list") (acceptingei: "'state list") definition nbaei_rel where [to_relAPP]: "nbaei_rel L S \ {(A\<^sub>1, A\<^sub>2). (alphabetei A\<^sub>1, alphabetei A\<^sub>2) \ \L\ list_rel \ (initialei A\<^sub>1, initialei A\<^sub>2) \ \S\ list_rel \ (transitionei A\<^sub>1, transitionei A\<^sub>2) \ \S \\<^sub>r L \\<^sub>r S\ list_rel \ (acceptingei A\<^sub>1, acceptingei A\<^sub>2) \ \S\ list_rel}" lemma nbaei_param[param, autoref_rules]: "(nbaei, nbaei) \ \L\ list_rel \ \S\ list_rel \ \S \\<^sub>r L \\<^sub>r S\ list_rel \ \S\ list_rel \ \L, S\ nbaei_rel" "(alphabetei, alphabetei) \ \L, S\ nbaei_rel \ \L\ list_rel" "(initialei, initialei) \ \L, S\ nbaei_rel \ \S\ list_rel" "(transitionei, transitionei) \ \L, S\ nbaei_rel \ \S \\<^sub>r L \\<^sub>r S\ list_rel" "(acceptingei, acceptingei) \ \L, S\ nbaei_rel \ \S\ list_rel" unfolding nbaei_rel_def by auto definition nbaei_nbae_rel where [to_relAPP]: "nbaei_nbae_rel L S \ {(A\<^sub>1, A\<^sub>2). (alphabetei A\<^sub>1, alphabete A\<^sub>2) \ \L\ list_set_rel \ (initialei A\<^sub>1, initiale A\<^sub>2) \ \S\ list_set_rel \ (transitionei A\<^sub>1, transitione A\<^sub>2) \ \S \\<^sub>r L \\<^sub>r S\ list_set_rel \ (acceptingei A\<^sub>1, acceptinge A\<^sub>2) \ \S\ list_set_rel}" lemmas [autoref_rel_intf] = REL_INTFI[of nbaei_nbae_rel i_nbae_scheme] lemma nbaei_nbae_param[param, autoref_rules]: "(nbaei, nbae) \ \L\ list_set_rel \ \S\ list_set_rel \ \S \\<^sub>r L \\<^sub>r S\ list_set_rel \ \S\ list_set_rel \ \L, S\ nbaei_nbae_rel" "(alphabetei, alphabete) \ \L, S\ nbaei_nbae_rel \ \L\ list_set_rel" "(initialei, initiale) \ \L, S\ nbaei_nbae_rel \ \S\ list_set_rel" "(transitionei, transitione) \ \L, S\ nbaei_nbae_rel \ \S \\<^sub>r L \\<^sub>r S\ list_set_rel" "(acceptingei, acceptinge) \ \L, S\ nbaei_nbae_rel \ \S\ list_set_rel" unfolding nbaei_nbae_rel_def by auto definition nbaei_nbae where "nbaei_nbae A \ nbae (set (alphabetei A)) (set (initialei A)) (set (transitionei A)) (set (acceptingei A))" lemma nbaei_nbae_id_param[param]: "(nbaei_nbae, id) \ \L, S\ nbaei_nbae_rel \ \L, S\ nbae_rel" proof fix Ai A assume 1: "(Ai, A) \ \L, S\ nbaei_nbae_rel" have 2: "nbaei_nbae Ai = nbae (set (alphabetei Ai)) (set (initialei Ai)) (set (transitionei Ai)) (set (acceptingei Ai))" unfolding nbaei_nbae_def by rule have 3: "id A = nbae (id (alphabete A)) (id (initiale A)) (id (transitione A)) (id (acceptinge A))" by simp show "(nbaei_nbae Ai, id A) \ \L, S\ nbae_rel" unfolding 2 3 using 1 by parametricity qed abbreviation "transitions L S s \ \ a \ L. \ p \ S. {p} \ {a} \ s a p" abbreviation "succs T a p \ (T `` {p}) `` {a}" definition nba_nbae where "nba_nbae A \ nbae (alphabet A) (initial A) (transitions (alphabet A) (nodes A) (transition A)) (Set.filter (accepting A) (nodes A))" definition nbae_nba where "nbae_nba A \ nba (alphabete A) (initiale A) (succs (transitione A)) (\ p. p \ acceptinge A)" lemma nba_nbae_param[param]: "(nba_nbae, nba_nbae) \ \L, S\ nba_rel \ \L, S\ nbae_rel" unfolding nba_nbae_def by parametricity lemma nbae_nba_param[param]: assumes "bijective L" "bijective S" shows "(nbae_nba, nbae_nba) \ \L, S\ nbae_rel \ \L, S\ nba_rel" using assms assms(2)[unfolded bijective_alt] unfolding nbae_nba_def by parametricity auto lemma nbae_nba_nba_nbae_param[param]: "((nbae_nba \ nba_nbae) A, id A) \ \Id_on (alphabet A), Id_on (nodes A)\ nba_rel" proof - have "(nbae_nba \ nba_nbae) A = nba (alphabet A) (initial A) (succs (transitions (alphabet A) (nodes A) (transition A))) (\ p. p \ Set.filter (accepting A) (nodes A))" unfolding nbae_nba_def nba_nbae_def by simp also have "(\, nba (alphabet A) (initial A) (transition A) (accepting A)) \ \Id_on (alphabet A), Id_on (nodes A)\ nba_rel" using nba_rel_eq by parametricity auto also have "nba (alphabet A) (initial A) (transition A) (accepting A) = id A" by simp finally show ?thesis by this qed definition nbaei_nba_rel where [to_relAPP]: "nbaei_nba_rel L S \ {(Ae, A). (nbae_nba (nbaei_nbae Ae), A) \ \L, S\ nba_rel}" lemma nbaei_nba_id[param]: "(nbae_nba \ nbaei_nbae, id) \ \L, S\ nbaei_nba_rel \ \L, S\ nba_rel" unfolding nbaei_nba_rel_def by auto - schematic_goal nbae_nba_impl: "(?f, nbae_nba) \ \Id, Id\ nbaei_nbae_rel \ \Id, Id\ nbai_nba_rel" + schematic_goal nbae_nba_impl: + assumes [autoref_rules]: "(leq, HOL.eq) \ L \ L \ bool_rel" + assumes [autoref_rules]: "(seq, HOL.eq) \ S \ S \ bool_rel" + shows "(?f, nbae_nba) \ \L, S\ nbaei_nbae_rel \ \L, S\ nbai_nba_rel" unfolding nbae_nba_def by autoref concrete_definition nbae_nba_impl uses nbae_nba_impl + lemma nbae_nba_impl_refine[autoref_rules]: + assumes "GEN_OP leq HOL.eq (L \ L \ bool_rel)" + assumes "GEN_OP seq HOL.eq (S \ S \ bool_rel)" + shows "(nbae_nba_impl leq seq, nbae_nba) \ \L, S\ nbaei_nbae_rel \ \L, S\ nbai_nba_rel" + using nbae_nba_impl.refine assms unfolding autoref_tag_defs by this end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Implement.thy b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Implement.thy --- a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Implement.thy +++ b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Implement.thy @@ -1,82 +1,103 @@ section \Implementation of Nondeterministic Büchi Automata\ theory NBA_Implement imports "NBA_Refine" "../../Basic/Implement" begin + consts i_nba_scheme :: "interface \ interface \ interface" + + context + begin + + interpretation autoref_syn by this + + lemma nba_scheme_itype[autoref_itype]: + "nba ::\<^sub>i \L\\<^sub>i i_set \\<^sub>i \S\\<^sub>i i_set \\<^sub>i (L \\<^sub>i S \\<^sub>i \S\\<^sub>i i_set) \\<^sub>i \S\\<^sub>i i_set \\<^sub>i + \L, S\\<^sub>i i_nba_scheme" + "alphabet ::\<^sub>i \L, S\\<^sub>i i_nba_scheme \\<^sub>i \L\\<^sub>i i_set" + "initial ::\<^sub>i \L, S\\<^sub>i i_nba_scheme \\<^sub>i \S\\<^sub>i i_set" + "transition ::\<^sub>i \L, S\\<^sub>i i_nba_scheme \\<^sub>i L \\<^sub>i S \\<^sub>i \S\\<^sub>i i_set" + "accepting ::\<^sub>i \L, S\\<^sub>i i_nba_scheme \\<^sub>i \S\\<^sub>i i_set" + by auto + + end + datatype ('label, 'state) nbai = nbai (alphabeti: "'label list") (initiali: "'state list") (transitioni: "'label \ 'state \ 'state list") (acceptingi: "'state \ bool") definition nbai_rel :: "('label\<^sub>1 \ 'label\<^sub>2) set \ ('state\<^sub>1 \ 'state\<^sub>2) set \ (('label\<^sub>1, 'state\<^sub>1) nbai \ ('label\<^sub>2, 'state\<^sub>2) nbai) set" where [to_relAPP]: "nbai_rel L S \ {(A\<^sub>1, A\<^sub>2). (alphabeti A\<^sub>1, alphabeti A\<^sub>2) \ \L\ list_rel \ (initiali A\<^sub>1, initiali A\<^sub>2) \ \S\ list_rel \ (transitioni A\<^sub>1, transitioni A\<^sub>2) \ L \ S \ \S\ list_rel \ (acceptingi A\<^sub>1, acceptingi A\<^sub>2) \ S \ bool_rel}" - lemma nbai_param[param]: + lemma nbai_param[param, autoref_rules]: "(nbai, nbai) \ \L\ list_rel \ \S\ list_rel \ (L \ S \ \S\ list_rel) \ (S \ bool_rel) \ \L, S\ nbai_rel" "(alphabeti, alphabeti) \ \L, S\ nbai_rel \ \L\ list_rel" "(initiali, initiali) \ \L, S\ nbai_rel \ \S\ list_rel" "(transitioni, transitioni) \ \L, S\ nbai_rel \ L \ S \ \S\ list_rel" "(acceptingi, acceptingi) \ \L, S\ nbai_rel \ (S \ bool_rel)" unfolding nbai_rel_def fun_rel_def by auto definition nbai_nba_rel :: "('label\<^sub>1 \ 'label\<^sub>2) set \ ('state\<^sub>1 \ 'state\<^sub>2) set \ (('label\<^sub>1, 'state\<^sub>1) nbai \ ('label\<^sub>2, 'state\<^sub>2) nba) set" where [to_relAPP]: "nbai_nba_rel L S \ {(A\<^sub>1, A\<^sub>2). (alphabeti A\<^sub>1, alphabet A\<^sub>2) \ \L\ list_set_rel \ (initiali A\<^sub>1, initial A\<^sub>2) \ \S\ list_set_rel \ (transitioni A\<^sub>1, transition A\<^sub>2) \ L \ S \ \S\ list_set_rel \ (acceptingi A\<^sub>1, accepting A\<^sub>2) \ S \ bool_rel}" + lemmas [autoref_rel_intf] = REL_INTFI[of nbai_nba_rel i_nba_scheme] + + (* TODO: why is there a warning? *) lemma nbai_nba_param[param, autoref_rules]: "(nbai, nba) \ \L\ list_set_rel \ \S\ list_set_rel \ (L \ S \ \S\ list_set_rel) \ (S \ bool_rel) \ \L, S\ nbai_nba_rel" "(alphabeti, alphabet) \ \L, S\ nbai_nba_rel \ \L\ list_set_rel" "(initiali, initial) \ \L, S\ nbai_nba_rel \ \S\ list_set_rel" "(transitioni, transition) \ \L, S\ nbai_nba_rel \ L \ S \ \S\ list_set_rel" "(acceptingi, accepting) \ \L, S\ nbai_nba_rel \ S \ bool_rel" unfolding nbai_nba_rel_def fun_rel_def by auto definition nbai_nba :: "('label, 'state) nbai \ ('label, 'state) nba" where "nbai_nba A \ nba (set (alphabeti A)) (set (initiali A)) (\ a p. set (transitioni A a p)) (acceptingi A)" definition nbai_invar :: "('label, 'state) nbai \ bool" where "nbai_invar A \ distinct (alphabeti A) \ distinct (initiali A) \ (\ a p. distinct (transitioni A a p))" lemma nbai_nba_id_param[param]: "(nbai_nba, id) \ \L, S\ nbai_nba_rel \ \L, S\ nba_rel" proof fix Ai A assume 1: "(Ai, A) \ \L, S\ nbai_nba_rel" have 2: "nbai_nba Ai = nba (set (alphabeti Ai)) (set (initiali Ai)) (\ a p. set (transitioni Ai a p)) (acceptingi Ai)" unfolding nbai_nba_def by rule have 3: "id A = nba (id (alphabet A)) (id (initial A)) (\ a p. id (transition A a p)) (accepting A)" by simp show "(nbai_nba Ai, id A) \ \L, S\ nba_rel" unfolding 2 3 using 1 by parametricity qed lemma nbai_nba_br: "\Id, Id\ nbai_nba_rel = br nbai_nba nbai_invar" proof safe show "(A, B) \ \Id, Id\ nbai_nba_rel" if "(A, B) \ br nbai_nba nbai_invar" for A and B :: "('a, 'b) nba" using that unfolding nbai_nba_rel_def nbai_nba_def nbai_invar_def by (auto simp: in_br_conv list_set_rel_def) show "(A, B) \ br nbai_nba nbai_invar" if "(A, B) \ \Id, Id\ nbai_nba_rel" for A and B :: "('a, 'b) nba" proof - have 1: "(nbai_nba A, id B) \ \Id, Id\ nba_rel" using that by parametricity have 2: "nbai_invar A" using nbai_nba_param(2 - 5)[param_fo, OF that] by (auto simp: in_br_conv list_set_rel_def nbai_invar_def) show ?thesis using 1 2 unfolding in_br_conv by auto qed qed end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Translate.thy b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Translate.thy --- a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Translate.thy +++ b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Translate.thy @@ -1,253 +1,267 @@ section \Explore and Enumerate Nodes of Nondeterministic Büchi Automata\ theory NBA_Translate imports NBA_Explicit begin subsection \Syntax\ (* TODO: this syntax has unnecessarily high inner binding strength, requiring extra parentheses the regular let syntax correctly uses inner binding strength 0: ("(2_ =/ _)" 10) *) no_syntax "_do_let" :: "[pttrn, 'a] \ do_bind" ("(2let _ =/ _)" [1000, 13] 13) syntax "_do_let" :: "[pttrn, 'a] \ do_bind" ("(2let _ =/ _)" 13) section \Image on Explicit Automata\ definition nbae_image where "nbae_image f A \ nbae (alphabete A) (f ` initiale A) ((\ (p, a, q). (f p, a, f q)) ` transitione A) (f ` acceptinge A)" lemma nbae_image_param[param]: "(nbae_image, nbae_image) \ (S \ T) \ \L, S\ nbae_rel \ \L, T\ nbae_rel" unfolding nbae_image_def by parametricity lemma nbae_image_id[simp]: "nbae_image id = id" unfolding nbae_image_def by auto lemma nbae_image_nba_nbae: "nbae_image f (nba_nbae A) = nbae (alphabet A) (f ` initial A) (\ p \ nodes A. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p) (f ` {p \ nodes A. accepting A p})" unfolding nba_nbae_def nbae_image_def nbae.simps Set.filter_def by force section \Exploration and Translation\ definition trans_spec where "trans_spec A f \ \ p \ nodes A. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p" definition trans_algo where "trans_algo N L S f \ FOREACH N (\ p T. do { ASSERT (p \ N); FOREACH L (\ a T. do { ASSERT (a \ L); FOREACH (S a p) (\ q T. do { ASSERT (q \ S a p); ASSERT ((f p, a, f q) \ T); RETURN (insert (f p, a, f q) T) } ) T } ) T } ) {}" lemma trans_algo_refine: assumes "finite (nodes A)" "finite (alphabet A)" "inj_on f (nodes A)" assumes "N = nodes A" "L = alphabet A" "S = transition A" shows "(trans_algo N L S f, SPEC (HOL.eq (trans_spec A f))) \ \Id\ nres_rel" unfolding trans_algo_def trans_spec_def assms(4-6) proof (refine_vcg FOREACH_rule_insert_eq) show "finite (nodes A)" using assms(1) by this show "(\ p \ nodes A. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p) = (\ p \ nodes A. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p)" by rule show "(\ p \ {}. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p) = {}" by simp fix T x assume 1: "T \ nodes A" "x \ nodes A" "x \ T" show "finite (alphabet A)" using assms(2) by this show "(\ a \ {}. f ` {x} \ {a} \ f ` transition A a x) \ (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p) = (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p)" "(\ a \ alphabet A. f ` {x} \ {a} \ f ` transition A a x) \ (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p) = (\ p \ insert x T. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p)" by auto fix Ta xa assume 2: "Ta \ alphabet A" "xa \ alphabet A" "xa \ Ta" show "finite (transition A xa x)" using 1 2 assms(1) by (meson infinite_subset nba.nodes_transition subsetI) show "(f ` {x} \ {xa} \ f ` transition A xa x) \ (\ a \ Ta. f ` {x} \ {a} \ f ` transition A a x) \ (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p) = (\ a \ insert xa Ta. f ` {x} \ {a} \ f ` transition A a x) \ (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p)" by auto show "(f ` {x} \ {xa} \ f ` {}) \ (\ a \ Ta. f ` {x} \ {a} \ f ` transition A a x) \ (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p) = (\ a \ Ta. f ` {x} \ {a} \ f ` transition A a x) \ (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p)" by auto fix Tb xb assume 3: "Tb \ transition A xa x" "xb \ transition A xa x" "xb \ Tb" show "(f x, xa, f xb) \ f ` {x} \ {xa} \ f ` Tb \ (\ a \ Ta. f ` {x} \ {a} \ f ` transition A a x) \ (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p)" using 1 2 3 assms(3) by (blast dest: inj_onD) show "f ` {x} \ {xa} \ f ` insert xb Tb \ (\ a \ Ta. f ` {x} \ {a} \ f ` transition A a x) \ (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p) = insert (f x, xa, f xb) (f ` {x} \ {xa} \ f ` Tb \ (\ a \ Ta. f ` {x} \ {a} \ f ` transition A a x) \ (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p))" by auto qed + (* TODO: this operation is fundamentally nondeterministic, model it as such, instead of + putting everything in the relation *) + (* TODO: by setting this to id, we pretend that nothing is actually happening with this + translation and put all the weight on the relation, it would make more sense to model + the fact that a translation is happening explicitly *) definition to_nbaei :: "('state, 'label) nba \ ('state, 'label) nba" where "to_nbaei \ id" - (* TODO: make separate implementations for "nba_nbae" and "op_set_enumerate \ nbae_image" *) + (* TODO: make separate implementations for "nba_nbae" and "op_set_enumerate \ nbae_image" + make sure to do regression tests along the way *) schematic_goal to_nbaei_impl: fixes S :: "('statei \ 'state) set" assumes [simp]: "finite (nodes A)" assumes [autoref_ga_rules]: "is_bounded_hashcode S seq bhc" assumes [autoref_ga_rules]: "is_valid_def_hm_size TYPE('statei) hms" assumes [autoref_rules]: "(seq, HOL.eq) \ S \ S \ bool_rel" assumes [autoref_rules]: "(Ai, A) \ \L, S\ nbai_nba_rel" shows "(?f :: ?'a, do { let N = nodes A; f \ op_set_enumerate N; ASSERT (dom f = N); ASSERT (\ p \ initial A. f p \ None); ASSERT (\ a \ alphabet A. \ p \ dom f. \ q \ transition A a p. f q \ None); T \ trans_algo N (alphabet A) (transition A) (\ x. the (f x)); RETURN (nbae (alphabet A) ((\ x. the (f x)) ` initial A) T ((\ x. the (f x)) ` {p \ N. accepting A p})) }) \ ?R" unfolding trans_algo_def by (autoref_monadic (plain)) concrete_definition to_nbaei_impl uses to_nbaei_impl lemma to_nbaei_impl_refine'': fixes S :: "('statei \ 'state) set" assumes "finite (nodes A)" assumes "is_bounded_hashcode S seq bhc" assumes "is_valid_def_hm_size TYPE('statei) hms" assumes "(seq, HOL.eq) \ S \ S \ bool_rel" assumes "(Ai, A) \ \L, S\ nbai_nba_rel" shows "(RETURN (to_nbaei_impl seq bhc hms Ai), do { f \ op_set_enumerate (nodes A); RETURN (nbae_image (the \ f) (nba_nbae A)) }) \ \\L, nat_rel\ nbaei_nbae_rel\ nres_rel" proof - have 1: "finite (alphabet A)" using nbai_nba_param(2)[param_fo, OF assms(5)] list_set_rel_finite unfolding finite_set_rel_def by auto note to_nbaei_impl.refine[OF assms] also have "(do { let N = nodes A; f \ op_set_enumerate N; ASSERT (dom f = N); ASSERT (\ p \ initial A. f p \ None); ASSERT (\ a \ alphabet A. \ p \ dom f. \ q \ transition A a p. f q \ None); T \ trans_algo N (alphabet A) (transition A) (\ x. the (f x)); RETURN (nbae (alphabet A) ((\ x. the (f x)) ` initial A) T ((\ x. the (f x)) ` {p \ N. accepting A p})) }, do { f \ op_set_enumerate (nodes A); T \ SPEC (HOL.eq (trans_spec A (\ x. the (f x)))); RETURN (nbae (alphabet A) ((\ x. the (f x)) ` initial A) T ((\ x. the (f x)) ` {p \ nodes A. accepting A p})) }) \ \Id\ nres_rel" unfolding Let_def comp_apply op_set_enumerate_def using assms(1) 1 by (refine_vcg vcg0[OF trans_algo_refine]) (auto intro!: inj_on_map_the[unfolded comp_apply]) also have "(do { f \ op_set_enumerate (nodes A); T \ SPEC (HOL.eq (trans_spec A (\ x. the (f x)))); RETURN (nbae (alphabet A) ((\ x. the (f x)) ` initial A) T ((\ x. the (f x)) ` {p \ nodes A. accepting A p})) }, do { f \ op_set_enumerate (nodes A); RETURN (nbae_image (the \ f) (nba_nbae A)) }) \ \Id\ nres_rel" unfolding trans_spec_def nbae_image_nba_nbae by refine_vcg force finally show ?thesis unfolding nres_rel_comp by simp qed (* TODO: generalize L *) context fixes Ai A fixes seq bhc hms fixes S :: "('statei \ 'state) set" assumes a: "finite (nodes A)" assumes b: "is_bounded_hashcode S seq bhc" assumes c: "is_valid_def_hm_size TYPE('statei) hms" assumes d: "(seq, HOL.eq) \ S \ S \ bool_rel" assumes e: "(Ai, A) \ \Id, S\ nbai_nba_rel" begin definition f' where "f' \ SOME f'. (to_nbaei_impl seq bhc hms Ai, nbae_image (the \ f') (nba_nbae A)) \ \Id, nat_rel\ nbaei_nbae_rel \ dom f' = nodes A \ inj_on f' (nodes A)" lemma 1: "\ f'. (to_nbaei_impl seq bhc hms Ai, nbae_image (the \ f') (nba_nbae A)) \ \Id, nat_rel\ nbaei_nbae_rel \ dom f' = nodes A \ inj_on f' (nodes A)" using to_nbaei_impl_refine''[ OF a b c d e, unfolded op_set_enumerate_def bind_RES_RETURN_eq, THEN nres_relD, THEN RETURN_ref_SPECD] by force lemma f'_refine: "(to_nbaei_impl seq bhc hms Ai, nbae_image (the \ f') (nba_nbae A)) \ \Id, nat_rel\ nbaei_nbae_rel" using someI_ex[OF 1, folded f'_def] by auto lemma f'_dom: "dom f' = nodes A" using someI_ex[OF 1, folded f'_def] by auto lemma f'_inj: "inj_on f' (nodes A)" using someI_ex[OF 1, folded f'_def] by auto definition f where "f \ the \ f'" definition g where "g = inv_into (nodes A) f" lemma inj_f[intro!, simp]: "inj_on f (nodes A)" using f'_inj f'_dom unfolding f_def by (simp add: inj_on_map_the) lemma inj_g[intro!, simp]: "inj_on g (f ` nodes A)" unfolding g_def by (simp add: inj_on_inv_into) definition rel where "rel \ {(f p, p) |p. p \ nodes A}" lemma rel_alt_def: "rel = (br f (\ p. p \ nodes A))\" unfolding rel_def by (auto simp: in_br_conv) lemma rel_inv_def: "rel = br g (\ k. k \ f ` nodes A)" unfolding rel_alt_def g_def by (auto simp: in_br_conv) lemma rel_domain[simp]: "Domain rel = f ` nodes A" unfolding rel_def by force lemma rel_range[simp]: "Range rel = nodes A" unfolding rel_def by auto - lemma [intro!, simp]: "bijective rel" unfolding rel_inv_def by (simp add: bijective_alt) + lemma rel_bijective[intro!, simp]: "bijective rel" unfolding rel_inv_def by (simp add: bijective_alt) lemma [simp]: "Id_on (f ` nodes A) O rel = rel" unfolding rel_def by auto lemma [simp]: "rel O Id_on (nodes A) = rel" unfolding rel_def by auto lemma [param]: "(f, f) \ Id_on (Range rel) \ Id_on (Domain rel)" unfolding rel_alt_def by auto lemma [param]: "(g, g) \ Id_on (Domain rel) \ Id_on (Range rel)" unfolding rel_inv_def by auto lemma [param]: "(id, f) \ rel \ Id_on (Domain rel)" unfolding rel_alt_def by (auto simp: in_br_conv) lemma [param]: "(f, id) \ Id_on (Range rel) \ rel" unfolding rel_alt_def by (auto simp: in_br_conv) lemma [param]: "(id, g) \ Id_on (Domain rel) \ rel" unfolding rel_inv_def by (auto simp: in_br_conv) lemma [param]: "(g, id) \ rel \ Id_on (Range rel)" unfolding rel_inv_def by (auto simp: in_br_conv) + lemma f_refine: "(to_nbaei_impl seq bhc hms Ai, nbae_image f (nba_nbae A)) \ + \Id, nat_rel\ nbaei_nbae_rel" using f'_refine unfolding f_def by this + + lemma nba_f_refine: "(nbae_nba (nbae_image f (nba_nbae A)), A) \ \Id_on (alphabet A), rel\ nba_rel" + proof - + have "(nbae_nba (nbae_image f (nba_nbae A)), nbae_nba (nbae_image id (nba_nbae A))) \ + \Id_on (alphabet A), rel\ nba_rel" using nba_rel_eq by parametricity auto + also have "nbae_nba (nbae_image id (nba_nbae A)) = (nbae_nba \ nba_nbae) A" by simp + also have "(\, id A) \ \Id_on (alphabet A), Id_on (nodes A)\ nba_rel" by parametricity + finally show ?thesis by simp + qed lemma to_nbaei_impl_refine': "(to_nbaei_impl seq bhc hms Ai, to_nbaei A) \ \Id_on (alphabet A), rel\ nbaei_nba_rel" proof - have "(nbae_nba (nbaei_nbae (to_nbaei_impl seq bhc hms Ai)), nbae_nba (id (nbae_image f (nba_nbae A)))) \ - \Id, nat_rel\ nba_rel" using f'_refine[folded f_def] by parametricity auto - also have "(nbae_nba (id (nbae_image f (nba_nbae A))), nbae_nba (id (nbae_image id (nba_nbae A)))) \ - \Id_on (alphabet A), rel\ nba_rel" using nba_rel_eq by parametricity auto - also have "nbae_nba (id (nbae_image id (nba_nbae A))) = (nbae_nba \ nba_nbae) A" by simp - also have "(\, id A) \ \Id_on (alphabet A), Id_on (nodes A)\ nba_rel" by parametricity - also have "id A = to_nbaei A" unfolding to_nbaei_def by simp - finally show ?thesis unfolding nbaei_nba_rel_def by simp + \Id, nat_rel\ nba_rel" using f_refine by parametricity auto + also have "(nbae_nba (id (nbae_image f (nba_nbae A))), A) \ + \Id_on (alphabet A), rel\ nba_rel" using nba_f_refine by simp + finally show ?thesis unfolding nbaei_nba_rel_def to_nbaei_def by simp qed end context begin interpretation autoref_syn by this lemma to_nbaei_impl_refine[autoref_rules]: fixes S :: "('statei \ 'state) set" assumes "SIDE_PRECOND (finite (nodes A))" assumes "SIDE_GEN_ALGO (is_bounded_hashcode S seq bhc)" assumes "SIDE_GEN_ALGO (is_valid_def_hm_size TYPE('statei) hms)" assumes "GEN_OP seq HOL.eq (S \ S \ bool_rel)" assumes "(Ai, A) \ \Id, S\ nbai_nba_rel" shows "(to_nbaei_impl seq bhc hms Ai, (OP to_nbaei ::: \Id, S\ nbai_nba_rel \ \Id_on (alphabet A), rel Ai A seq bhc hms\ nbaei_nba_rel) $ A) \ \Id_on (alphabet A), rel Ai A seq bhc hms\ nbaei_nba_rel" using to_nbaei_impl_refine' assms unfolding autoref_tag_defs by this end end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/NBA/NGBA_Algorithms.thy b/thys/Transition_Systems_and_Automata/Automata/NBA/NGBA_Algorithms.thy new file mode 100644 --- /dev/null +++ b/thys/Transition_Systems_and_Automata/Automata/NBA/NGBA_Algorithms.thy @@ -0,0 +1,67 @@ +section \Algorithms on Nondeterministic Generalized Büchi Automata\ + +theory NGBA_Algorithms +imports + NBA_Algorithms + NGBA_Implement + NBA_Combine +begin + + subsection \Implementations\ + + context + begin + + interpretation autoref_syn by this + + (* TODO: move *) + lemma degen_param[param, autoref_rules]: "(degen, degen) \ \S \ bool_rel\ list_rel \ S \\<^sub>r nat_rel \ bool_rel" + proof (intro fun_relI) + fix cs ds ak bl + assume "(cs, ds) \ \S \ bool_rel\ list_rel" "(ak, bl) \ S \\<^sub>r nat_rel" + then show "(degen cs ak, degen ds bl) \ bool_rel" + unfolding degen_def list_rel_def fun_rel_def list_all2_conv_all_nth + by (cases "snd ak < length cs") (auto 0 3) + qed + + lemma count_alt_def: "Degeneralization.count cs a k = + (if k < length cs + then if (cs ! k) a then Suc k mod length cs else k + else if List.null cs then k else 0)" + unfolding count_def null_def by rule + + (* TODO: move *) + lemmas [param] = null_transfer[unfolded pred_bool_Id, to_set] + + (* TODO: move *) + lemma count_param[param, autoref_rules]: "(Degeneralization.count, Degeneralization.count) \ + \A \ bool_rel\ list_rel \ A \ nat_rel \ nat_rel" + unfolding count_alt_def by parametricity + + (* TODO: why do we need this? *) + lemma degeneralize_annotated: "degeneralize A = nba + (ngba.alphabet A) + (ngba.initial A \ ({0} ::: \nat_rel\ list_set_rel)) + (\ a (p, k). ngba.transition A a p \ ({Degeneralization.count (ngba.accepting A) p k} ::: \nat_rel\ list_set_rel)) + (degen (ngba.accepting A))" + unfolding degeneralization.degeneralize_def by simp + + schematic_goal ngba_degeneralize: "(?f :: ?'a, degeneralize) \ ?R" + unfolding degeneralize_annotated by autoref + concrete_definition ngba_degeneralize uses ngba_degeneralize + lemmas ngba_degeneralize_refine[autoref_rules] = ngba_degeneralize.refine + + schematic_goal nba_intersect': + assumes [autoref_rules]: "(seq, HOL.eq) \ L \ L \ bool_rel" + shows "(?f, intersect') \ \L, S\ nbai_nba_rel \ \L, T\ nbai_nba_rel \ \L, S \\<^sub>r T\ ngbai_ngba_rel" + unfolding intersection.intersect_def by autoref + concrete_definition nba_intersect' uses nba_intersect' + lemma nba_intersect'_refine[autoref_rules]: + assumes "GEN_OP seq HOL.eq (L \ L \ bool_rel)" + shows "(nba_intersect' seq, intersect') \ + \L, S\ nbai_nba_rel \ \L, T\ nbai_nba_rel \ \L, S \\<^sub>r T\ ngbai_ngba_rel" + using nba_intersect'.refine assms unfolding autoref_tag_defs by this + + end + +end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/NBA/NGBA_Implement.thy b/thys/Transition_Systems_and_Automata/Automata/NBA/NGBA_Implement.thy new file mode 100644 --- /dev/null +++ b/thys/Transition_Systems_and_Automata/Automata/NBA/NGBA_Implement.thy @@ -0,0 +1,102 @@ +section \Implementation of Nondeterministic Generalized Büchi Automata\ + +theory NGBA_Implement +imports + "NGBA_Refine" + "../../Basic/Implement" +begin + + consts i_ngba_scheme :: "interface \ interface \ interface" + + context + begin + + interpretation autoref_syn by this + + lemma ngba_scheme_itype[autoref_itype]: + "ngba ::\<^sub>i \L\\<^sub>i i_set \\<^sub>i \S\\<^sub>i i_set \\<^sub>i (L \\<^sub>i S \\<^sub>i \S\\<^sub>i i_set) \\<^sub>i \\S\\<^sub>i i_set\\<^sub>i i_list \\<^sub>i + \L, S\\<^sub>i i_ngba_scheme" + "alphabet ::\<^sub>i \L, S\\<^sub>i i_ngba_scheme \\<^sub>i \L\\<^sub>i i_set" + "initial ::\<^sub>i \L, S\\<^sub>i i_ngba_scheme \\<^sub>i \S\\<^sub>i i_set" + "transition ::\<^sub>i \L, S\\<^sub>i i_ngba_scheme \\<^sub>i L \\<^sub>i S \\<^sub>i \S\\<^sub>i i_set" + "accepting ::\<^sub>i \L, S\\<^sub>i i_ngba_scheme \\<^sub>i \\S\\<^sub>i i_set\\<^sub>i i_list" + by auto + + end + + datatype ('label, 'state) ngbai = ngbai + (alphabeti: "'label list") + (initiali: "'state list") + (transitioni: "'label \ 'state \ 'state list") + (acceptingi: "('state \ bool) list") + + definition ngbai_rel :: "('label\<^sub>1 \ 'label\<^sub>2) set \ ('state\<^sub>1 \ 'state\<^sub>2) set \ + (('label\<^sub>1, 'state\<^sub>1) ngbai \ ('label\<^sub>2, 'state\<^sub>2) ngbai) set" where + [to_relAPP]: "ngbai_rel L S \ {(A\<^sub>1, A\<^sub>2). + (alphabeti A\<^sub>1, alphabeti A\<^sub>2) \ \L\ list_rel \ + (initiali A\<^sub>1, initiali A\<^sub>2) \ \S\ list_rel \ + (transitioni A\<^sub>1, transitioni A\<^sub>2) \ L \ S \ \S\ list_rel \ + (acceptingi A\<^sub>1, acceptingi A\<^sub>2) \ \S \ bool_rel\ list_rel}" + + lemma ngbai_param[param]: + "(ngbai, ngbai) \ \L\ list_rel \ \S\ list_rel \ (L \ S \ \S\ list_rel) \ + \S \ bool_rel\ list_rel \ \L, S\ ngbai_rel" + "(alphabeti, alphabeti) \ \L, S\ ngbai_rel \ \L\ list_rel" + "(initiali, initiali) \ \L, S\ ngbai_rel \ \S\ list_rel" + "(transitioni, transitioni) \ \L, S\ ngbai_rel \ L \ S \ \S\ list_rel" + "(acceptingi, acceptingi) \ \L, S\ ngbai_rel \ \S \ bool_rel\ list_rel" + unfolding ngbai_rel_def fun_rel_def by auto + + definition ngbai_ngba_rel :: "('label\<^sub>1 \ 'label\<^sub>2) set \ ('state\<^sub>1 \ 'state\<^sub>2) set \ + (('label\<^sub>1, 'state\<^sub>1) ngbai \ ('label\<^sub>2, 'state\<^sub>2) ngba) set" where + [to_relAPP]: "ngbai_ngba_rel L S \ {(A\<^sub>1, A\<^sub>2). + (alphabeti A\<^sub>1, alphabet A\<^sub>2) \ \L\ list_set_rel \ + (initiali A\<^sub>1, initial A\<^sub>2) \ \S\ list_set_rel \ + (transitioni A\<^sub>1, transition A\<^sub>2) \ L \ S \ \S\ list_set_rel \ + (acceptingi A\<^sub>1, accepting A\<^sub>2) \ \S \ bool_rel\ list_rel}" + + lemmas [autoref_rel_intf] = REL_INTFI[of ngbai_ngba_rel i_ngba_scheme] + + lemma ngbai_ngba_param[param, autoref_rules]: + "(ngbai, ngba) \ \L\ list_set_rel \ \S\ list_set_rel \ (L \ S \ \S\ list_set_rel) \ + \S \ bool_rel\ list_rel \ \L, S\ ngbai_ngba_rel" + "(alphabeti, alphabet) \ \L, S\ ngbai_ngba_rel \ \L\ list_set_rel" + "(initiali, initial) \ \L, S\ ngbai_ngba_rel \ \S\ list_set_rel" + "(transitioni, transition) \ \L, S\ ngbai_ngba_rel \ L \ S \ \S\ list_set_rel" + "(acceptingi, accepting) \ \L, S\ ngbai_ngba_rel \ \S \ bool_rel\ list_rel" + unfolding ngbai_ngba_rel_def fun_rel_def by auto + + definition ngbai_ngba :: "('label, 'state) ngbai \ ('label, 'state) ngba" where + "ngbai_ngba A \ ngba (set (alphabeti A)) (set (initiali A)) (\ a p. set (transitioni A a p)) (acceptingi A)" + definition ngbai_invar :: "('label, 'state) ngbai \ bool" where + "ngbai_invar A \ distinct (alphabeti A) \ distinct (initiali A) \ (\ a p. distinct (transitioni A a p))" + + lemma ngbai_ngba_id_param[param]: "(ngbai_ngba, id) \ \L, S\ ngbai_ngba_rel \ \L, S\ ngba_rel" + proof + fix Ai A + assume 1: "(Ai, A) \ \L, S\ ngbai_ngba_rel" + have 2: "ngbai_ngba Ai = ngba (set (alphabeti Ai)) (set (initiali Ai)) + (\ a p. set (transitioni Ai a p)) (acceptingi Ai)" unfolding ngbai_ngba_def by rule + have 3: "id A = ngba (id (alphabet A)) (id (initial A)) + (\ a p. id (transition A a p)) (accepting A)" by simp + show "(ngbai_ngba Ai, id A) \ \L, S\ ngba_rel" unfolding 2 3 using 1 by parametricity + qed + + lemma ngbai_ngba_br: "\Id, Id\ ngbai_ngba_rel = br ngbai_ngba ngbai_invar" + proof safe + show "(A, B) \ \Id, Id\ ngbai_ngba_rel" if "(A, B) \ br ngbai_ngba ngbai_invar" + for A and B :: "('a, 'b) ngba" + using that unfolding ngbai_ngba_rel_def ngbai_ngba_def ngbai_invar_def + by (auto simp: in_br_conv list_set_rel_def) + show "(A, B) \ br ngbai_ngba ngbai_invar" if "(A, B) \ \Id, Id\ ngbai_ngba_rel" + for A and B :: "('a, 'b) ngba" + proof - + have 1: "(ngbai_ngba A, id B) \ \Id, Id\ ngba_rel" using that by parametricity + have 2: "ngbai_invar A" + using ngbai_ngba_param(2 - 5)[param_fo, OF that] + by (auto simp: in_br_conv list_set_rel_def ngbai_invar_def) + show ?thesis using 1 2 unfolding in_br_conv by auto + qed + qed + +end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/NBA/NGBA_Refine.thy b/thys/Transition_Systems_and_Automata/Automata/NBA/NGBA_Refine.thy new file mode 100644 --- /dev/null +++ b/thys/Transition_Systems_and_Automata/Automata/NBA/NGBA_Refine.thy @@ -0,0 +1,57 @@ +section \Relations on Nondeterministic Generalized Büchi Automata\ + +theory NGBA_Refine +imports + "NGBA" + "../../Transition_Systems/Transition_System_Refine" +begin + + definition ngba_rel :: "('label\<^sub>1 \ 'label\<^sub>2) set \ ('state\<^sub>1 \ 'state\<^sub>2) set \ + (('label\<^sub>1, 'state\<^sub>1) ngba \ ('label\<^sub>2, 'state\<^sub>2) ngba) set" where + [to_relAPP]: "ngba_rel L S \ {(A\<^sub>1, A\<^sub>2). + (alphabet A\<^sub>1, alphabet A\<^sub>2) \ \L\ set_rel \ + (initial A\<^sub>1, initial A\<^sub>2) \ \S\ set_rel \ + (transition A\<^sub>1, transition A\<^sub>2) \ L \ S \ \S\ set_rel \ + (accepting A\<^sub>1, accepting A\<^sub>2) \ \S \ bool_rel\ list_rel}" + + lemma ngba_param[param]: + "(ngba, ngba) \ \L\ set_rel \ \S\ set_rel \ (L \ S \ \S\ set_rel) \ \S \ bool_rel\ list_rel \ + \L, S\ ngba_rel" + "(alphabet, alphabet) \ \L, S\ ngba_rel \ \L\ set_rel" + "(initial, initial) \ \L, S\ ngba_rel \ \S\ set_rel" + "(transition, transition) \ \L, S\ ngba_rel \ L \ S \ \S\ set_rel" + "(accepting, accepting) \ \L, S\ ngba_rel \ \S \ bool_rel\ list_rel" + unfolding ngba_rel_def fun_rel_def by auto + + lemma ngba_rel_id[simp]: "\Id, Id\ ngba_rel = Id" unfolding ngba_rel_def using ngba.expand by auto + + lemma enableds_param[param]: "(ngba.enableds, ngba.enableds) \ \L, S\ ngba_rel \ S \ \L \\<^sub>r S\ set_rel" + using ngba_param(2, 4) unfolding ngba.enableds_def fun_rel_def set_rel_def by fastforce + lemma paths_param[param]: "(ngba.paths, ngba.paths) \ \L, S\ ngba_rel \ S \ \\L \\<^sub>r S\ list_rel\ set_rel" + using enableds_param[param_fo] by parametricity + lemma runs_param[param]: "(ngba.runs, ngba.runs) \ \L, S\ ngba_rel \ S \ \\L \\<^sub>r S\ stream_rel\ set_rel" + using enableds_param[param_fo] by parametricity + + lemma reachable_param[param]: "(reachable, reachable) \ \L, S\ ngba_rel \ S \ \S\ set_rel" + proof - + have 1: "reachable A p = (\ wr. target wr p) ` ngba.paths A p" for A :: "('label, 'state) ngba" and p + unfolding ngba.reachable_alt_def ngba.paths_def by auto + show ?thesis unfolding 1 using enableds_param[param_fo] by parametricity + qed + lemma nodes_param[param]: "(nodes, nodes) \ \L, S\ ngba_rel \ \S\ set_rel" + unfolding ngba.nodes_alt_def Collect_mem_eq by parametricity + + (* TODO: move *) + lemma gen_param[param]: "(gen, gen) \ (A \ B \ bool_rel) \ \A\ list_rel \ B \ bool_rel" + unfolding gen_def by parametricity + + lemma language_param[param]: "(language, language) \ \L, S\ ngba_rel \ \\L\ stream_rel\ set_rel" + proof - + have 1: "language A = (\ p \ initial A. \ wr \ ngba.runs A p. + if gen infs (accepting A) (trace wr p) then {smap fst wr} else {})" + for A :: "('label, 'state) ngba" + unfolding ngba.language_def ngba.runs_def image_def by (auto iff: split_szip_ex) + show ?thesis unfolding 1 using enableds_param[param_fo] by parametricity + qed + +end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/Nondeterministic.thy b/thys/Transition_Systems_and_Automata/Automata/Nondeterministic.thy --- a/thys/Transition_Systems_and_Automata/Automata/Nondeterministic.thy +++ b/thys/Transition_Systems_and_Automata/Automata/Nondeterministic.thy @@ -1,392 +1,438 @@ theory Nondeterministic imports "../Transition_Systems/Transition_System" "../Transition_Systems/Transition_System_Extra" "../Transition_Systems/Transition_System_Construction" "../Basic/Degeneralization" begin type_synonym ('label, 'state) trans = "'label \ 'state \ 'state set" locale automaton = fixes automaton :: "'label set \ 'state set \ ('label, 'state) trans \ 'condition \ 'automaton" fixes alphabet :: "'automaton \ 'label set" fixes initial :: "'automaton \ 'state set" fixes transition :: "'automaton \ ('label, 'state) trans" fixes condition :: "'automaton \ 'condition" assumes automaton[simp]: "automaton (alphabet A) (initial A) (transition A) (condition A) = A" assumes alphabet[simp]: "alphabet (automaton a i t c) = a" assumes initial[simp]: "initial (automaton a i t c) = i" assumes transition[simp]: "transition (automaton a i t c) = t" assumes condition[simp]: "condition (automaton a i t c) = c" begin sublocale transition_system_initial "\ a p. snd a" "\ a p. fst a \ alphabet A \ snd a \ transition A (fst a) p" "\ p. p \ initial A" for A defines path' = path and run' = run and reachable' = reachable and nodes' = nodes by this lemma states_alt_def: "states r p = map snd r" by (induct r arbitrary: p) (auto) lemma trace_alt_def: "trace r p = smap snd r" by (coinduction arbitrary: r p) (auto) lemma successors_alt_def: "successors A p = (\ a \ alphabet A. transition A a p)" by auto lemma reachable_transition[intro]: assumes "a \ alphabet A" "q \ reachable A p" "r \ transition A a q" shows "r \ reachable A p" using reachable.execute assms by force lemma nodes_transition[intro]: assumes "a \ alphabet A" "p \ nodes A" "q \ transition A a p" shows "q \ nodes A" using nodes.execute assms by force definition restrict :: "'automaton \ 'automaton" where "restrict A \ automaton (alphabet A) (initial A) - (\ a. if a \ alphabet A then transition A a else bot) + (\ a p. if a \ alphabet A then transition A a p else {}) (condition A)" lemma restrict_simps[simp]: "alphabet (restrict A) = alphabet A" "initial (restrict A) = initial A" "transition (restrict A) a p = (if a \ alphabet A then transition A a p else {})" "condition (restrict A) = condition A" unfolding restrict_def by auto lemma restrict_path[simp]: "path (restrict A) = path A" proof (intro ext iffI) show "path A wr p" if "path (restrict A) wr p" for wr p using that by induct auto show "path (restrict A) wr p" if "path A wr p" for wr p using that by induct auto qed lemma restrict_run[simp]: "run (restrict A) = run A" proof (intro ext iffI) show "run A wr p" if "run (restrict A) wr p" for wr p using that by coinduct auto show "run (restrict A) wr p" if "run A wr p" for wr p using that by coinduct auto qed end (* TODO: create analogous thing for NFAs (automaton_target) *) locale automaton_trace = automaton automaton alphabet initial transition condition for automaton :: "'label set \ 'state set \ ('label, 'state) trans \ 'condition \ 'automaton" and alphabet :: "'automaton \ 'label set" and initial :: "'automaton \ 'state set" and transition :: "'automaton \ ('label, 'state) trans" and condition :: "'automaton \ 'condition" + fixes test :: "'condition \ 'state stream \ bool" begin definition language :: "'automaton \ 'label stream set" where "language A \ {w |w r p. p \ initial A \ run A (w ||| r) p \ test (condition A) (trace (w ||| r) p)}" lemma language[intro]: assumes "p \ initial A" "run A (w ||| r) p" "test (condition A) (trace (w ||| r) p)" shows "w \ language A" using assms unfolding language_def by auto lemma language_elim[elim]: assumes "w \ language A" obtains r p where "p \ initial A" "run A (w ||| r) p" "test (condition A) (trace (w ||| r) p)" using assms unfolding language_def by auto lemma run_alphabet: assumes "run A (w ||| r) p" shows "w \ streams (alphabet A)" using assms by (coinduction arbitrary: w r p) (metis run.cases stream.map szip_smap szip_smap_fst) lemma language_alphabet: "language A \ streams (alphabet A)" unfolding language_def by (auto dest: run_alphabet) lemma restrict_language[simp]: "language (restrict A) = language A" by force end locale automaton_degeneralization = a: automaton automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 + b: automaton automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 for automaton\<^sub>1 :: "'label set \ 'state set \ ('label, 'state) trans \ 'state pred gen \ 'automaton\<^sub>1" and alphabet\<^sub>1 :: "'automaton\<^sub>1 \ 'label set" and initial\<^sub>1 :: "'automaton\<^sub>1 \ 'state set" and transition\<^sub>1 :: "'automaton\<^sub>1 \ ('label, 'state) trans" and condition\<^sub>1 :: "'automaton\<^sub>1 \ 'state pred gen" and automaton\<^sub>2 :: "'label set \ 'state degen set \ ('label, 'state degen) trans \ 'state degen pred \ 'automaton\<^sub>2" and alphabet\<^sub>2 :: "'automaton\<^sub>2 \ 'label set" and initial\<^sub>2 :: "'automaton\<^sub>2 \ 'state degen set" and transition\<^sub>2 :: "'automaton\<^sub>2 \ ('label, 'state degen) trans" and condition\<^sub>2 :: "'automaton\<^sub>2 \ 'state degen pred" begin definition degeneralize :: "'automaton\<^sub>1 \ 'automaton\<^sub>2" where "degeneralize A \ automaton\<^sub>2 (alphabet\<^sub>1 A) (initial\<^sub>1 A \ {0}) (\ a (p, k). transition\<^sub>1 A a p \ {count (condition\<^sub>1 A) p k}) (degen (condition\<^sub>1 A))" lemma degeneralize_simps[simp]: "alphabet\<^sub>2 (degeneralize A) = alphabet\<^sub>1 A" "initial\<^sub>2 (degeneralize A) = initial\<^sub>1 A \ {0}" "transition\<^sub>2 (degeneralize A) a (p, k) = transition\<^sub>1 A a p \ {count (condition\<^sub>1 A) p k}" "condition\<^sub>2 (degeneralize A) = degen (condition\<^sub>1 A)" unfolding degeneralize_def by auto lemma run_degeneralize: assumes "a.run A (w ||| r) p" shows "b.run (degeneralize A) (w ||| r ||| sscan (count (condition\<^sub>1 A)) (p ## r) k) (p, k)" using assms by (coinduction arbitrary: w r p k) (force elim: a.run.cases) lemma degeneralize_run: assumes "b.run (degeneralize A) (w ||| rs) pk" obtains r s p k where "rs = r ||| s" "pk = (p, k)" "a.run A (w ||| r) p" "s = sscan (count (condition\<^sub>1 A)) (p ## r) k" proof show "rs = smap fst rs ||| smap snd rs" "pk = (fst pk, snd pk)" by auto show "a.run A (w ||| smap fst rs) (fst pk)" using assms by (coinduction arbitrary: w rs pk) (force elim: b.run.cases) show "smap snd rs = sscan (count (condition\<^sub>1 A)) (fst pk ## smap fst rs) (snd pk)" using assms by (coinduction arbitrary: w rs pk) (force elim: b.run.cases) qed + lemma degeneralize_nodes: + "b.nodes (degeneralize A) \ a.nodes A \ insert 0 {0 ..< length (condition\<^sub>1 A)}" + proof + fix pk + assume "pk \ b.nodes (degeneralize A)" + then show "pk \ a.nodes A \ insert 0 {0 ..< length (condition\<^sub>1 A)}" + by (induct) (force, cases "condition\<^sub>1 A = []", auto) + qed + lemma nodes_degeneralize: "a.nodes A \ fst ` b.nodes (degeneralize A)" + proof + fix p + assume "p \ a.nodes A" + then show "p \ fst ` b.nodes (degeneralize A)" + proof induct + case (initial p) + have "(p, 0) \ b.nodes (degeneralize A)" using initial by auto + then show ?case using image_iff fst_conv by force + next + case (execute p aq) + obtain k where "(p, k) \ b.nodes (degeneralize A)" using execute(2) by auto + then have "(snd aq, count (condition\<^sub>1 A) p k) \ b.nodes (degeneralize A)" + using execute(3) by auto + then show ?case using image_iff snd_conv by force + qed + qed + + lemma degeneralize_nodes_finite[iff]: "finite (b.nodes (degeneralize A)) \ finite (a.nodes A)" + proof + show "finite (a.nodes A)" if "finite (b.nodes (degeneralize A))" + using finite_subset nodes_degeneralize that by blast + show "finite (b.nodes (degeneralize A))" if "finite (a.nodes A)" + using finite_subset degeneralize_nodes that by blast + qed + end locale automaton_degeneralization_trace = automaton_degeneralization automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 + a: automaton_trace automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_trace automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 for automaton\<^sub>1 :: "'label set \ 'state set \ ('label, 'state) trans \ 'state pred gen \ 'automaton\<^sub>1" and alphabet\<^sub>1 :: "'automaton\<^sub>1 \ 'label set" and initial\<^sub>1 :: "'automaton\<^sub>1 \ 'state set" and transition\<^sub>1 :: "'automaton\<^sub>1 \ ('label, 'state) trans" and condition\<^sub>1 :: "'automaton\<^sub>1 \ 'state pred gen" and test\<^sub>1 :: "'state pred gen \ 'state stream \ bool" and automaton\<^sub>2 :: "'label set \ 'state degen set \ ('label, 'state degen) trans \ 'state degen pred \ 'automaton\<^sub>2" and alphabet\<^sub>2 :: "'automaton\<^sub>2 \ 'label set" and initial\<^sub>2 :: "'automaton\<^sub>2 \ 'state degen set" and transition\<^sub>2 :: "'automaton\<^sub>2 \ ('label, 'state degen) trans" and condition\<^sub>2 :: "'automaton\<^sub>2 \ 'state degen pred" and test\<^sub>2 :: "'state degen pred \ 'state degen stream \ bool" + assumes test[iff]: "test\<^sub>2 (degen cs) (r ||| sscan (count cs) (p ## r) k) \ test\<^sub>1 cs r" begin lemma degeneralize_language[simp]: "b.language (degeneralize A) = a.language A" unfolding a.language_def b.language_def unfolding a.trace_alt_def b.trace_alt_def by (auto dest: run_degeneralize elim!: degeneralize_run) end locale automaton_intersection = a: automaton automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 + b: automaton automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 + c: automaton automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 for automaton\<^sub>1 :: "'label set \ 'state\<^sub>1 set \ ('label, 'state\<^sub>1) trans \ 'condition\<^sub>1 \ 'automaton\<^sub>1" and alphabet\<^sub>1 :: "'automaton\<^sub>1 \ 'label set" and initial\<^sub>1 :: "'automaton\<^sub>1 \ 'state\<^sub>1 set" and transition\<^sub>1 :: "'automaton\<^sub>1 \ ('label, 'state\<^sub>1) trans" and condition\<^sub>1 :: "'automaton\<^sub>1 \ 'condition\<^sub>1" and automaton\<^sub>2 :: "'label set \ 'state\<^sub>2 set \ ('label, 'state\<^sub>2) trans \ 'condition\<^sub>2 \ 'automaton\<^sub>2" and alphabet\<^sub>2 :: "'automaton\<^sub>2 \ 'label set" and initial\<^sub>2 :: "'automaton\<^sub>2 \ 'state\<^sub>2 set" and transition\<^sub>2 :: "'automaton\<^sub>2 \ ('label, 'state\<^sub>2) trans" and condition\<^sub>2 :: "'automaton\<^sub>2 \ 'condition\<^sub>2" and automaton\<^sub>3 :: "'label set \ ('state\<^sub>1 \ 'state\<^sub>2) set \ ('label, 'state\<^sub>1 \ 'state\<^sub>2) trans \ 'condition\<^sub>3 \ 'automaton\<^sub>3" and alphabet\<^sub>3 :: "'automaton\<^sub>3 \ 'label set" and initial\<^sub>3 :: "'automaton\<^sub>3 \ ('state\<^sub>1 \ 'state\<^sub>2) set" and transition\<^sub>3 :: "'automaton\<^sub>3 \ ('label, 'state\<^sub>1 \ 'state\<^sub>2) trans" and condition\<^sub>3 :: "'automaton\<^sub>3 \ 'condition\<^sub>3" + fixes condition :: "'condition\<^sub>1 \ 'condition\<^sub>2 \ 'condition\<^sub>3" begin definition intersect :: "'automaton\<^sub>1 \ 'automaton\<^sub>2 \ 'automaton\<^sub>3" where "intersect A B \ automaton\<^sub>3 (alphabet\<^sub>1 A \ alphabet\<^sub>2 B) (initial\<^sub>1 A \ initial\<^sub>2 B) (\ a (p, q). transition\<^sub>1 A a p \ transition\<^sub>2 B a q) (condition (condition\<^sub>1 A) (condition\<^sub>2 B))" lemma intersect_simps[simp]: "alphabet\<^sub>3 (intersect A B) = alphabet\<^sub>1 A \ alphabet\<^sub>2 B" "initial\<^sub>3 (intersect A B) = initial\<^sub>1 A \ initial\<^sub>2 B" "transition\<^sub>3 (intersect A B) a (p, q) = transition\<^sub>1 A a p \ transition\<^sub>2 B a q" "condition\<^sub>3 (intersect A B) = condition (condition\<^sub>1 A) (condition\<^sub>2 B)" unfolding intersect_def by auto lemma intersect_path[iff]: assumes "length w = length r" "length r = length s" shows "c.path (intersect A B) (w || r || s) (p, q) \ a.path A (w || r) p \ b.path B (w || s) q" using assms by (induct arbitrary: p q rule: list_induct3) (auto) lemma intersect_run[iff]: "c.run (intersect A B) (w ||| r ||| s) (p, q) \ a.run A (w ||| r) p \ b.run B (w ||| s) q" proof safe show "a.run A (w ||| r) p" if "c.run (intersect A B) (w ||| r ||| s) (p, q)" using that by (coinduction arbitrary: w r s p q) (force elim: c.run.cases) show "b.run B (w ||| s) q" if "c.run (intersect A B) (w ||| r ||| s) (p, q)" using that by (coinduction arbitrary: w r s p q) (force elim: c.run.cases) show "c.run (intersect A B) (w ||| r ||| s) (p, q)" if "a.run A (w ||| r) p" "b.run B (w ||| s) q" using that by (coinduction arbitrary: w r s p q) (auto elim: a.run.cases b.run.cases) qed + lemma intersect_nodes: "c.nodes (intersect A B) \ a.nodes A \ b.nodes B" + proof + fix pq + assume "pq \ c.nodes (intersect A B)" + then show "pq \ a.nodes A \ b.nodes B" by induct auto + qed + + lemma intersect_nodes_finite[intro]: + assumes "finite (a.nodes A)" "finite (b.nodes B)" + shows "finite (c.nodes (intersect A B))" + using finite_subset intersect_nodes assms by blast + end locale automaton_intersection_trace = automaton_intersection automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 condition + a: automaton_trace automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_trace automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 + c: automaton_trace automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 test\<^sub>3 for automaton\<^sub>1 :: "'label set \ 'state\<^sub>1 set \ ('label, 'state\<^sub>1) trans \ 'condition\<^sub>1 \ 'automaton\<^sub>1" and alphabet\<^sub>1 :: "'automaton\<^sub>1 \ 'label set" and initial\<^sub>1 :: "'automaton\<^sub>1 \ 'state\<^sub>1 set" and transition\<^sub>1 :: "'automaton\<^sub>1 \ ('label, 'state\<^sub>1) trans" and condition\<^sub>1 :: "'automaton\<^sub>1 \ 'condition\<^sub>1" and test\<^sub>1 :: "'condition\<^sub>1 \ 'state\<^sub>1 stream \ bool" and automaton\<^sub>2 :: "'label set \ 'state\<^sub>2 set \ ('label, 'state\<^sub>2) trans \ 'condition\<^sub>2 \ 'automaton\<^sub>2" and alphabet\<^sub>2 :: "'automaton\<^sub>2 \ 'label set" and initial\<^sub>2 :: "'automaton\<^sub>2 \ 'state\<^sub>2 set" and transition\<^sub>2 :: "'automaton\<^sub>2 \ ('label, 'state\<^sub>2) trans" and condition\<^sub>2 :: "'automaton\<^sub>2 \ 'condition\<^sub>2" and test\<^sub>2 :: "'condition\<^sub>2 \ 'state\<^sub>2 stream \ bool" and automaton\<^sub>3 :: "'label set \ ('state\<^sub>1 \ 'state\<^sub>2) set \ ('label, 'state\<^sub>1 \ 'state\<^sub>2) trans \ 'condition\<^sub>3 \ 'automaton\<^sub>3" and alphabet\<^sub>3 :: "'automaton\<^sub>3 \ 'label set" and initial\<^sub>3 :: "'automaton\<^sub>3 \ ('state\<^sub>1 \ 'state\<^sub>2) set" and transition\<^sub>3 :: "'automaton\<^sub>3 \ ('label, 'state\<^sub>1 \ 'state\<^sub>2) trans" and condition\<^sub>3 :: "'automaton\<^sub>3 \ 'condition\<^sub>3" and test\<^sub>3 :: "'condition\<^sub>3 \ ('state\<^sub>1 \ 'state\<^sub>2) stream \ bool" and condition :: "'condition\<^sub>1 \ 'condition\<^sub>2 \ 'condition\<^sub>3" + assumes test[iff]: "test\<^sub>3 (condition c\<^sub>1 c\<^sub>2) (u ||| v) \ test\<^sub>1 c\<^sub>1 u \ test\<^sub>2 c\<^sub>2 v" begin - lemma combine_language[simp]: "c.language (intersect A B) = a.language A \ b.language B" + lemma intersect_language[simp]: "c.language (intersect A B) = a.language A \ b.language B" unfolding a.language_def b.language_def c.language_def unfolding a.trace_alt_def b.trace_alt_def c.trace_alt_def by (fastforce iff: split_szip) end locale automaton_union = a: automaton automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 + b: automaton automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 + c: automaton automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 for automaton\<^sub>1 :: "'label set \ 'state\<^sub>1 set \ ('label, 'state\<^sub>1) trans \ 'condition\<^sub>1 \ 'automaton\<^sub>1" and alphabet\<^sub>1 :: "'automaton\<^sub>1 \ 'label set" and initial\<^sub>1 :: "'automaton\<^sub>1 \ 'state\<^sub>1 set" and transition\<^sub>1 :: "'automaton\<^sub>1 \ ('label, 'state\<^sub>1) trans" and condition\<^sub>1 :: "'automaton\<^sub>1 \ 'condition\<^sub>1" and automaton\<^sub>2 :: "'label set \ 'state\<^sub>2 set \ ('label, 'state\<^sub>2) trans \ 'condition\<^sub>2 \ 'automaton\<^sub>2" and alphabet\<^sub>2 :: "'automaton\<^sub>2 \ 'label set" and initial\<^sub>2 :: "'automaton\<^sub>2 \ 'state\<^sub>2 set" and transition\<^sub>2 :: "'automaton\<^sub>2 \ ('label, 'state\<^sub>2) trans" and condition\<^sub>2 :: "'automaton\<^sub>2 \ 'condition\<^sub>2" and automaton\<^sub>3 :: "'label set \ ('state\<^sub>1 + 'state\<^sub>2) set \ ('label, 'state\<^sub>1 + 'state\<^sub>2) trans \ 'condition\<^sub>3 \ 'automaton\<^sub>3" and alphabet\<^sub>3 :: "'automaton\<^sub>3 \ 'label set" and initial\<^sub>3 :: "'automaton\<^sub>3 \ ('state\<^sub>1 + 'state\<^sub>2) set" and transition\<^sub>3 :: "'automaton\<^sub>3 \ ('label, 'state\<^sub>1 + 'state\<^sub>2) trans" and condition\<^sub>3 :: "'automaton\<^sub>3 \ 'condition\<^sub>3" + fixes condition :: "'condition\<^sub>1 \ 'condition\<^sub>2 \ 'condition\<^sub>3" begin definition union :: "'automaton\<^sub>1 \ 'automaton\<^sub>2 \ 'automaton\<^sub>3" where "union A B \ automaton\<^sub>3 (alphabet\<^sub>1 A \ alphabet\<^sub>2 B) (initial\<^sub>1 A <+> initial\<^sub>2 B) (\ a. \ Inl p \ Inl ` transition\<^sub>1 A a p | Inr q \ Inr ` transition\<^sub>2 B a q) (condition (condition\<^sub>1 A) (condition\<^sub>2 B))" lemma union_simps[simp]: "alphabet\<^sub>3 (union A B) = alphabet\<^sub>1 A \ alphabet\<^sub>2 B" "initial\<^sub>3 (union A B) = initial\<^sub>1 A <+> initial\<^sub>2 B" "transition\<^sub>3 (union A B) a (Inl p) = Inl ` transition\<^sub>1 A a p" "transition\<^sub>3 (union A B) a (Inr q) = Inr ` transition\<^sub>2 B a q" "condition\<^sub>3 (union A B) = condition (condition\<^sub>1 A) (condition\<^sub>2 B)" unfolding union_def by auto lemma run_union_a: assumes "a.run A (w ||| r) p" shows "c.run (union A B) (w ||| smap Inl r) (Inl p)" using assms by (coinduction arbitrary: w r p) (force elim: a.run.cases) lemma run_union_b: assumes "b.run B (w ||| s) q" shows "c.run (union A B) (w ||| smap Inr s) (Inr q)" using assms by (coinduction arbitrary: w s q) (force elim: b.run.cases) lemma run_union: assumes "alphabet\<^sub>1 A = alphabet\<^sub>2 B" assumes "c.run (union A B) (w ||| rs) pq" obtains (a) r p where "rs = smap Inl r" "pq = Inl p" "a.run A (w ||| r) p" | (b) s q where "rs = smap Inr s" "pq = Inr q" "b.run B (w ||| s) q" proof (cases pq) case (Inl p) have 1: "rs = smap Inl (smap projl rs)" using assms(2) unfolding Inl by (coinduction arbitrary: w rs p) (force elim: c.run.cases) have 2: "a.run A (w ||| smap projl rs) p" using assms unfolding Inl by (coinduction arbitrary: w rs p) (force elim: c.run.cases) show ?thesis using a 1 Inl 2 by this next case (Inr q) have 1: "rs = smap Inr (smap projr rs)" using assms(2) unfolding Inr by (coinduction arbitrary: w rs q) (force elim: c.run.cases) have 2: "b.run B (w ||| smap projr rs) q" using assms unfolding Inr by (coinduction arbitrary: w rs q) (force elim: c.run.cases) show ?thesis using b 1 Inr 2 by this qed end locale automaton_union_trace = automaton_union automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 condition + a: automaton_trace automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_trace automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 + c: automaton_trace automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 test\<^sub>3 for automaton\<^sub>1 :: "'label set \ 'state\<^sub>1 set \ ('label, 'state\<^sub>1) trans \ 'condition\<^sub>1 \ 'automaton\<^sub>1" and alphabet\<^sub>1 :: "'automaton\<^sub>1 \ 'label set" and initial\<^sub>1 :: "'automaton\<^sub>1 \ 'state\<^sub>1 set" and transition\<^sub>1 :: "'automaton\<^sub>1 \ ('label, 'state\<^sub>1) trans" and condition\<^sub>1 :: "'automaton\<^sub>1 \ 'condition\<^sub>1" and test\<^sub>1 :: "'condition\<^sub>1 \ 'state\<^sub>1 stream \ bool" and automaton\<^sub>2 :: "'label set \ 'state\<^sub>2 set \ ('label, 'state\<^sub>2) trans \ 'condition\<^sub>2 \ 'automaton\<^sub>2" and alphabet\<^sub>2 :: "'automaton\<^sub>2 \ 'label set" and initial\<^sub>2 :: "'automaton\<^sub>2 \ 'state\<^sub>2 set" and transition\<^sub>2 :: "'automaton\<^sub>2 \ ('label, 'state\<^sub>2) trans" and condition\<^sub>2 :: "'automaton\<^sub>2 \ 'condition\<^sub>2" and test\<^sub>2 :: "'condition\<^sub>2 \ 'state\<^sub>2 stream \ bool" and automaton\<^sub>3 :: "'label set \ ('state\<^sub>1 + 'state\<^sub>2) set \ ('label, 'state\<^sub>1 + 'state\<^sub>2) trans \ 'condition\<^sub>3 \ 'automaton\<^sub>3" and alphabet\<^sub>3 :: "'automaton\<^sub>3 \ 'label set" and initial\<^sub>3 :: "'automaton\<^sub>3 \ ('state\<^sub>1 + 'state\<^sub>2) set" and transition\<^sub>3 :: "'automaton\<^sub>3 \ ('label, 'state\<^sub>1 + 'state\<^sub>2) trans" and condition\<^sub>3 :: "'automaton\<^sub>3 \ 'condition\<^sub>3" and test\<^sub>3 :: "'condition\<^sub>3 \ ('state\<^sub>1 + 'state\<^sub>2) stream \ bool" and condition :: "'condition\<^sub>1 \ 'condition\<^sub>2 \ 'condition\<^sub>3" + assumes test\<^sub>1[iff]: "test\<^sub>3 (condition c\<^sub>1 c\<^sub>2) (smap Inl u) \ test\<^sub>1 c\<^sub>1 u" assumes test\<^sub>2[iff]: "test\<^sub>3 (condition c\<^sub>1 c\<^sub>2) (smap Inr v) \ test\<^sub>2 c\<^sub>2 v" begin lemma union_language[simp]: assumes "alphabet\<^sub>1 A = alphabet\<^sub>2 B" shows "c.language (union A B) = a.language A \ b.language B" using assms unfolding a.language_def b.language_def c.language_def unfolding a.trace_alt_def b.trace_alt_def c.trace_alt_def by (auto dest: run_union_a run_union_b elim!: run_union) end end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/ROOT b/thys/Transition_Systems_and_Automata/ROOT --- a/thys/Transition_Systems_and_Automata/ROOT +++ b/thys/Transition_Systems_and_Automata/ROOT @@ -1,36 +1,38 @@ chapter AFP session Transition_Systems_and_Automata (AFP) = "Collections" + options [timeout = 2400] sessions "DFS_Framework" "Gabow_SCC" directories + "Basic" + "Transition_Systems" "Automata" "Automata/DBA" "Automata/DCA" "Automata/DFA" "Automata/DRA" "Automata/NBA" "Automata/NFA" - "Basic" - "Transition_Systems" theories "Basic/Basic" "Basic/Sequence" "Basic/Sequence_Zip" "Basic/Sequence_LTL" "Basic/Maps" "Basic/Acceptance" "Transition_Systems/Transition_System" "Transition_Systems/Transition_System_Extra" "Transition_Systems/Transition_System_Construction" "Automata/DFA/DFA" "Automata/NFA/NFA" + "Automata/NBA/NBA_Combine" "Automata/NBA/NBA_Translate" + "Automata/NBA/NGBA_Algorithms" "Automata/DBA/DBA_Combine" "Automata/DCA/DCA_Combine" "Automata/DRA/DRA_Combine" "Automata/DRA/DRA_Translate" document_files "root.tex" diff --git a/thys/Word_Lib/Word_Lemmas.thy b/thys/Word_Lib/Word_Lemmas.thy --- a/thys/Word_Lib/Word_Lemmas.thy +++ b/thys/Word_Lib/Word_Lemmas.thy @@ -1,5728 +1,5728 @@ (* * Copyright 2014, NICTA * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * @TAG(NICTA_BSD) *) section "Lemmas with Generic Word Length" theory Word_Lemmas imports Complex_Main Word_Next Word_Enum "HOL-Library.Sublist" begin text \Set up quickcheck to support words\ quickcheck_generator word constructors: "zero_class.zero :: ('a::len) word", "numeral :: num \ ('a::len) word", "uminus :: ('a::len) word \ ('a::len) word" instantiation Enum.finite_1 :: len begin definition "len_of_finite_1 (x :: Enum.finite_1 itself) \ (1 :: nat)" instance by (standard, auto simp: len_of_finite_1_def) end instantiation Enum.finite_2 :: len begin definition "len_of_finite_2 (x :: Enum.finite_2 itself) \ (2 :: nat)" instance by (standard, auto simp: len_of_finite_2_def) end instantiation Enum.finite_3 :: len begin definition "len_of_finite_3 (x :: Enum.finite_3 itself) \ (4 :: nat)" instance by (standard, auto simp: len_of_finite_3_def) end (* Provide wf and less_induct for word. wf may be more useful in loop proofs, less_induct in recursion proofs. *) lemma word_less_wf: "wf {(a, b). a < (b :: ('a::len) word)}" apply (rule wf_subset) apply (rule wf_measure) apply safe apply (subst in_measure) apply (erule unat_mono) done lemma word_less_induct: "\ \x::('a::len) word. (\y. y < x \ P y) \ P x \ \ P a" using word_less_wf by induct blast instantiation word :: (len) wellorder begin instance by (intro_classes) (metis word_less_induct) end lemma word_plus_mono_left: fixes x :: "'a :: len word" shows "\y \ z; x \ x + z\ \ y + x \ z + x" by unat_arith lemma word_2p_mult_inc: assumes x: "2 * 2 ^ n < (2::'a::len word) * 2 ^ m" assumes suc_n: "Suc n < LENGTH('a::len)" assumes suc_m: "Suc m < LENGTH('a::len)" assumes 2: "unat (2::'a::len word) = 2" shows "2^n < (2::'a::len word)^m" proof - from suc_n have "(2::nat) * 2 ^ n mod 2 ^ LENGTH('a::len) = 2 * 2^n" apply (subst mod_less) apply (subst power_Suc[symmetric]) apply (rule power_strict_increasing) apply simp apply simp apply simp done moreover from suc_m have "(2::nat) * 2 ^ m mod 2 ^ LENGTH('a::len) = 2 * 2^m" apply (subst mod_less) apply (subst power_Suc[symmetric]) apply (rule power_strict_increasing) apply simp apply simp apply simp done ultimately have "2 * 2 ^ n < (2::nat) * 2 ^ m" using x apply (unfold word_less_nat_alt) apply simp apply (subst (asm) unat_word_ariths(2))+ apply (subst (asm) 2)+ apply (subst (asm) word_unat_power, subst (asm) unat_of_nat)+ apply (simp add: mod_mult_right_eq) done with suc_n suc_m show ?thesis unfolding word_less_nat_alt apply (subst word_unat_power, subst unat_of_nat)+ apply simp done qed lemma word_power_increasing: assumes x: "2 ^ x < (2 ^ y::'a::len word)" "x < LENGTH('a::len)" "y < LENGTH('a::len)" assumes 2: "unat (2::'a::len word) = 2" shows "x < y" using x apply (induct x arbitrary: y) apply (case_tac y; simp) apply (case_tac y; clarsimp) apply (subst (asm) power_Suc [symmetric]) apply (subst (asm) p2_eq_0) apply simp apply (drule (2) word_2p_mult_inc, rule 2) apply simp done lemma word_shiftl_add_distrib: fixes x :: "'a :: len word" shows "(x + y) << n = (x << n) + (y << n)" by (simp add: shiftl_t2n ring_distribs) lemma less_Suc_unat_less_bound: "n < Suc (unat (x :: 'a :: len word)) \ n < 2 ^ LENGTH('a)" by (auto elim!: order_less_le_trans intro: Suc_leI) lemma up_ucast_inj: "\ ucast x = (ucast y::'b::len word); LENGTH('a) \ len_of TYPE ('b) \ \ x = (y::'a::len word)" by (subst (asm) bang_eq) (fastforce simp: nth_ucast word_size intro: word_eqI) lemmas ucast_up_inj = up_ucast_inj lemma up_ucast_inj_eq: "LENGTH('a) \ len_of TYPE ('b) \ (ucast x = (ucast y::'b::len word)) = (x = (y::'a::len word))" by (fastforce dest: up_ucast_inj) lemma no_plus_overflow_neg: "(x :: 'a :: len word) < -y \ x \ x + y" apply (simp add: no_plus_overflow_uint_size word_less_alt uint_word_ariths word_size) apply (subst(asm) zmod_zminus1_eq_if) apply (simp split: if_split_asm) done lemma ucast_ucast_eq: fixes x :: "'a::len word" fixes y :: "'b::len word" shows "\ ucast x = (ucast (ucast y::'a::len word)::'c::len word); LENGTH('a) \ LENGTH('b); LENGTH('b) \ LENGTH('c) \ \ x = ucast y" by (fastforce intro: word_eqI simp: bang_eq nth_ucast word_size) text \right-padding a word to a certain length\ definition "bl_pad_to bl sz \ bl @ (replicate (sz - length bl) False)" lemma bl_pad_to_length: assumes lbl: "length bl \ sz" shows "length (bl_pad_to bl sz) = sz" using lbl by (simp add: bl_pad_to_def) lemma bl_pad_to_prefix: "prefix bl (bl_pad_to bl sz)" by (simp add: bl_pad_to_def) lemma same_length_is_parallel: assumes len: "\y \ set as. length y = x" shows "\x \ set as. \y \ set as - {x}. x \ y" proof (rule, rule) fix x y assume xi: "x \ set as" and yi: "y \ set as - {x}" from len obtain q where len': "\y \ set as. length y = q" .. show "x \ y" proof (rule not_equal_is_parallel) from xi yi show "x \ y" by auto from xi yi len' show "length x = length y" by (auto dest: bspec) qed qed text \Lemmas about words\ lemmas and_bang = word_and_nth lemma of_drop_to_bl: "of_bl (drop n (to_bl x)) = (x && mask (size x - n))" apply (clarsimp simp: bang_eq test_bit_of_bl rev_nth cong: rev_conj_cong) apply (safe, simp_all add: word_size to_bl_nth) done lemma word_add_offset_less: fixes x :: "'a :: len word" assumes yv: "y < 2 ^ n" and xv: "x < 2 ^ m" and mnv: "sz < LENGTH('a :: len)" and xv': "x < 2 ^ (LENGTH('a :: len) - n)" and mn: "sz = m + n" shows "x * 2 ^ n + y < 2 ^ sz" proof (subst mn) from mnv mn have nv: "n < LENGTH('a)" and mv: "m < LENGTH('a)" by auto have uy: "unat y < 2 ^ n" by (rule order_less_le_trans [OF unat_mono [OF yv] order_eq_refl], rule unat_power_lower[OF nv]) have ux: "unat x < 2 ^ m" by (rule order_less_le_trans [OF unat_mono [OF xv] order_eq_refl], rule unat_power_lower[OF mv]) then show "x * 2 ^ n + y < 2 ^ (m + n)" using ux uy nv mnv xv' apply (subst word_less_nat_alt) apply (subst unat_word_ariths)+ apply (subst mod_less) apply simp apply (subst mult.commute) apply (rule nat_less_power_trans [OF _ order_less_imp_le [OF nv]]) apply (rule order_less_le_trans [OF unat_mono [OF xv']]) apply (cases "n = 0") apply simp apply simp apply (subst unat_power_lower[OF nv]) apply (subst mod_less) apply (erule order_less_le_trans [OF nat_add_offset_less], assumption) apply (rule mn) apply simp apply (simp add: mn mnv) apply (erule nat_add_offset_less) apply simp+ done qed lemma word_less_power_trans: fixes n :: "'a :: len word" assumes nv: "n < 2 ^ (m - k)" and kv: "k \ m" and mv: "m < len_of TYPE ('a)" shows "2 ^ k * n < 2 ^ m" using nv kv mv apply - apply (subst word_less_nat_alt) apply (subst unat_word_ariths) apply (subst mod_less) apply simp apply (rule nat_less_power_trans) apply (erule order_less_trans [OF unat_mono]) apply simp apply simp apply simp apply (rule nat_less_power_trans) apply (subst unat_power_lower[where 'a = 'a, symmetric]) apply simp apply (erule unat_mono) apply simp done lemma word_less_sub_le[simp]: fixes x :: "'a :: len word" assumes nv: "n < LENGTH('a)" shows "(x \ 2 ^ n - 1) = (x < 2 ^ n)" proof - have "Suc (unat ((2::'a word) ^ n - 1)) = unat ((2::'a word) ^ n)" using nv by (metis Suc_pred' power_2_ge_iff unat_gt_0 unat_minus_one word_not_simps(1)) then show ?thesis using nv apply - apply (subst word_le_nat_alt) apply (subst less_Suc_eq_le [symmetric]) apply (erule ssubst) apply (subst word_less_nat_alt) apply (rule refl) done qed lemma Suc_unat_diff_1: fixes x :: "'a :: len word" assumes lt: "1 \ x" shows "Suc (unat (x - 1)) = unat x" proof - have "0 < unat x" by (rule order_less_le_trans [where y = 1], simp, subst unat_1 [symmetric], rule iffD1 [OF word_le_nat_alt lt]) then show ?thesis by ((subst unat_sub [OF lt])+, simp only: unat_1) qed lemma word_div_sub: fixes x :: "'a :: len word" assumes yx: "y \ x" and y0: "0 < y" shows "(x - y) div y = x div y - 1" apply (rule word_unat.Rep_eqD) apply (subst unat_div) apply (subst unat_sub [OF yx]) apply (subst unat_sub) apply (subst word_le_nat_alt) apply (subst unat_div) apply (subst le_div_geq) apply (rule order_le_less_trans [OF _ unat_mono [OF y0]]) apply simp apply (subst word_le_nat_alt [symmetric], rule yx) apply simp apply (subst unat_div) apply (subst le_div_geq [OF _ iffD1 [OF word_le_nat_alt yx]]) apply (rule order_le_less_trans [OF _ unat_mono [OF y0]]) apply simp apply simp done lemma word_mult_less_mono1: fixes i :: "'a :: len word" assumes ij: "i < j" and knz: "0 < k" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i * k < j * k" proof - from ij ujk knz have jk: "unat i * unat k < 2 ^ len_of TYPE ('a)" by (auto intro: order_less_subst2 simp: word_less_nat_alt elim: mult_less_mono1) then show ?thesis using ujk knz ij by (auto simp: word_less_nat_alt iffD1 [OF unat_mult_lem]) qed lemma word_mult_less_dest: fixes i :: "'a :: len word" assumes ij: "i * k < j * k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i < j" using uik ujk ij by (auto simp: word_less_nat_alt iffD1 [OF unat_mult_lem] elim: mult_less_mono1) lemma word_mult_less_cancel: fixes k :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k < j * k) = (i < j)" by (rule iffI [OF word_mult_less_dest [OF _ uik ujk] word_mult_less_mono1 [OF _ knz ujk]]) lemma Suc_div_unat_helper: assumes szv: "sz < LENGTH('a :: len)" and usszv: "us \ sz" shows "2 ^ (sz - us) = Suc (unat (((2::'a :: len word) ^ sz - 1) div 2 ^ us))" proof - note usv = order_le_less_trans [OF usszv szv] from usszv obtain q where qv: "sz = us + q" by (auto simp: le_iff_add) have "Suc (unat (((2:: 'a word) ^ sz - 1) div 2 ^ us)) = (2 ^ us + unat ((2:: 'a word) ^ sz - 1)) div 2 ^ us" apply (subst unat_div unat_power_lower[OF usv])+ apply (subst div_add_self1, simp+) done also have "\ = ((2 ^ us - 1) + 2 ^ sz) div 2 ^ us" using szv by (simp add: unat_minus_one) also have "\ = 2 ^ q + ((2 ^ us - 1) div 2 ^ us)" apply (subst qv) apply (subst power_add) apply (subst div_mult_self2; simp) done also have "\ = 2 ^ (sz - us)" using qv by simp finally show ?thesis .. qed lemma set_enum_word8_def: "(set enum::word8 set) = {0, 1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21, 22, 23, 24, 25, 26, 27, 28, 29, 30, 31, 32, 33, 34, 35, 36, 37, 38, 39, 40, 41, 42, 43, 44, 45, 46, 47, 48, 49, 50, 51, 52, 53, 54, 55, 56, 57, 58, 59, 60, 61, 62, 63, 64, 65, 66, 67, 68, 69, 70, 71, 72, 73, 74, 75, 76, 77, 78, 79, 80, 81, 82, 83, 84, 85, 86, 87, 88, 89, 90, 91, 92, 93, 94, 95, 96, 97, 98, 99, 100, 101, 102, 103, 104, 105, 106, 107, 108, 109, 110, 111, 112, 113, 114, 115, 116, 117, 118, 119, 120, 121, 122, 123, 124, 125, 126, 127, 128, 129, 130, 131, 132, 133, 134, 135, 136, 137, 138, 139, 140, 141, 142, 143, 144, 145, 146, 147, 148, 149, 150, 151, 152, 153, 154, 155, 156, 157, 158, 159, 160, 161, 162, 163, 164, 165, 166, 167, 168, 169, 170, 171, 172, 173, 174, 175, 176, 177, 178, 179, 180, 181, 182, 183, 184, 185, 186, 187, 188, 189, 190, 191, 192, 193, 194, 195, 196, 197, 198, 199, 200, 201, 202, 203, 204, 205, 206, 207, 208, 209, 210, 211, 212, 213, 214, 215, 216, 217, 218, 219, 220, 221, 222, 223, 224, 225, 226, 227, 228, 229, 230, 231, 232, 233, 234, 235, 236, 237, 238, 239, 240, 241, 242, 243, 244, 245, 246, 247, 248, 249, 250, 251, 252, 253, 254, 255}" by eval lemma set_strip_insert: "\ x \ insert a S; x \ a \ \ x \ S" by simp lemma word8_exhaust: fixes x :: word8 shows "\x \ 0; x \ 1; x \ 2; x \ 3; x \ 4; x \ 5; x \ 6; x \ 7; x \ 8; x \ 9; x \ 10; x \ 11; x \ 12; x \ 13; x \ 14; x \ 15; x \ 16; x \ 17; x \ 18; x \ 19; x \ 20; x \ 21; x \ 22; x \ 23; x \ 24; x \ 25; x \ 26; x \ 27; x \ 28; x \ 29; x \ 30; x \ 31; x \ 32; x \ 33; x \ 34; x \ 35; x \ 36; x \ 37; x \ 38; x \ 39; x \ 40; x \ 41; x \ 42; x \ 43; x \ 44; x \ 45; x \ 46; x \ 47; x \ 48; x \ 49; x \ 50; x \ 51; x \ 52; x \ 53; x \ 54; x \ 55; x \ 56; x \ 57; x \ 58; x \ 59; x \ 60; x \ 61; x \ 62; x \ 63; x \ 64; x \ 65; x \ 66; x \ 67; x \ 68; x \ 69; x \ 70; x \ 71; x \ 72; x \ 73; x \ 74; x \ 75; x \ 76; x \ 77; x \ 78; x \ 79; x \ 80; x \ 81; x \ 82; x \ 83; x \ 84; x \ 85; x \ 86; x \ 87; x \ 88; x \ 89; x \ 90; x \ 91; x \ 92; x \ 93; x \ 94; x \ 95; x \ 96; x \ 97; x \ 98; x \ 99; x \ 100; x \ 101; x \ 102; x \ 103; x \ 104; x \ 105; x \ 106; x \ 107; x \ 108; x \ 109; x \ 110; x \ 111; x \ 112; x \ 113; x \ 114; x \ 115; x \ 116; x \ 117; x \ 118; x \ 119; x \ 120; x \ 121; x \ 122; x \ 123; x \ 124; x \ 125; x \ 126; x \ 127; x \ 128; x \ 129; x \ 130; x \ 131; x \ 132; x \ 133; x \ 134; x \ 135; x \ 136; x \ 137; x \ 138; x \ 139; x \ 140; x \ 141; x \ 142; x \ 143; x \ 144; x \ 145; x \ 146; x \ 147; x \ 148; x \ 149; x \ 150; x \ 151; x \ 152; x \ 153; x \ 154; x \ 155; x \ 156; x \ 157; x \ 158; x \ 159; x \ 160; x \ 161; x \ 162; x \ 163; x \ 164; x \ 165; x \ 166; x \ 167; x \ 168; x \ 169; x \ 170; x \ 171; x \ 172; x \ 173; x \ 174; x \ 175; x \ 176; x \ 177; x \ 178; x \ 179; x \ 180; x \ 181; x \ 182; x \ 183; x \ 184; x \ 185; x \ 186; x \ 187; x \ 188; x \ 189; x \ 190; x \ 191; x \ 192; x \ 193; x \ 194; x \ 195; x \ 196; x \ 197; x \ 198; x \ 199; x \ 200; x \ 201; x \ 202; x \ 203; x \ 204; x \ 205; x \ 206; x \ 207; x \ 208; x \ 209; x \ 210; x \ 211; x \ 212; x \ 213; x \ 214; x \ 215; x \ 216; x \ 217; x \ 218; x \ 219; x \ 220; x \ 221; x \ 222; x \ 223; x \ 224; x \ 225; x \ 226; x \ 227; x \ 228; x \ 229; x \ 230; x \ 231; x \ 232; x \ 233; x \ 234; x \ 235; x \ 236; x \ 237; x \ 238; x \ 239; x \ 240; x \ 241; x \ 242; x \ 243; x \ 244; x \ 245; x \ 246; x \ 247; x \ 248; x \ 249; x \ 250; x \ 251; x \ 252; x \ 253; x \ 254; x \ 255\ \ P" apply (subgoal_tac "x \ set enum", subst (asm) set_enum_word8_def) apply (drule set_strip_insert, assumption)+ apply (erule emptyE) apply (subst enum_UNIV, rule UNIV_I) done lemma upto_enum_red': assumes lt: "1 \ X" shows "[(0::'a :: len word) .e. X - 1] = map of_nat [0 ..< unat X]" proof - have lt': "unat X < 2 ^ LENGTH('a)" by (rule unat_lt2p) show ?thesis apply (subst upto_enum_red) apply (simp del: upt.simps) apply (subst Suc_unat_diff_1 [OF lt]) apply (rule map_cong [OF refl]) apply (rule toEnum_of_nat) apply simp apply (erule order_less_trans [OF _ lt']) done qed lemma upto_enum_red2: assumes szv: "sz < LENGTH('a :: len)" shows "[(0:: 'a :: len word) .e. 2 ^ sz - 1] = map of_nat [0 ..< 2 ^ sz]" using szv apply (subst unat_power_lower[OF szv, symmetric]) apply (rule upto_enum_red') apply (subst word_le_nat_alt, simp) done lemma upto_enum_step_red: assumes szv: "sz < len_of (TYPE('a))" and usszv: "us \ sz" shows "[0 :: 'a :: len word , 2 ^ us .e. 2 ^ sz - 1] = map (\x. of_nat x * 2 ^ us) [0 ..< 2 ^ (sz - us)]" using szv unfolding upto_enum_step_def apply (subst if_not_P) apply (rule leD) apply (subst word_le_nat_alt) apply (subst unat_minus_one) apply simp apply simp apply simp apply (subst upto_enum_red) apply (simp del: upt.simps) apply (subst Suc_div_unat_helper [where 'a = 'a, OF szv usszv, symmetric]) apply clarsimp apply (subst toEnum_of_nat) apply (erule order_less_trans) using szv apply simp apply simp done lemma upto_enum_word: "[x .e. y] = map of_nat [unat x ..< Suc (unat y)]" apply (subst upto_enum_red) apply clarsimp apply (subst toEnum_of_nat) prefer 2 apply (rule refl) apply (erule disjE, simp) apply clarsimp apply (erule order_less_trans) apply simp done lemma word_upto_Cons_eq: "\x = z; x < y; Suc (unat y) < 2 ^ LENGTH('a)\ \ [x::'a::len word .e. y] = z # [x + 1 .e. y]" apply (subst upto_enum_red) apply (subst upt_conv_Cons) apply (simp) apply (drule unat_mono) apply arith apply (simp only: list.map) apply (subst list.inject) apply rule apply (rule to_from_enum) apply (subst upto_enum_red) apply (rule map_cong [OF _ refl]) apply (rule arg_cong2 [where f = "\x y. [x ..< y]"]) apply unat_arith apply simp done lemma distinct_enum_upto: "distinct [(0 :: 'a::len word) .e. b]" proof - have "\(b::'a word). [0 .e. b] = nths enum {..< Suc (fromEnum b)}" apply (subst upto_enum_red) apply (subst nths_upt_eq_take) apply (subst enum_word_def) apply (subst take_map) apply (subst take_upt) apply (simp only: add_0 fromEnum_unat) apply (rule order_trans [OF _ order_eq_refl]) apply (rule Suc_leI [OF unat_lt2p]) apply simp apply clarsimp apply (rule toEnum_of_nat) apply (erule order_less_trans [OF _ unat_lt2p]) done then show ?thesis by (rule ssubst) (rule distinct_nthsI, simp) qed lemma upto_enum_set_conv [simp]: fixes a :: "'a :: len word" shows "set [a .e. b] = {x. a \ x \ x \ b}" apply (subst upto_enum_red) apply (subst set_map) apply safe apply simp apply clarsimp apply (erule disjE) apply simp apply (erule iffD2 [OF word_le_nat_alt]) apply clarsimp apply (erule word_unat.Rep_cases [OF unat_le [OF order_less_imp_le]]) apply simp apply (erule iffD2 [OF word_le_nat_alt]) apply simp apply clarsimp apply (erule disjE) apply simp apply clarsimp apply (rule word_unat.Rep_cases [OF unat_le [OF order_less_imp_le]]) apply assumption apply simp apply (erule order_less_imp_le [OF iffD2 [OF word_less_nat_alt]]) apply clarsimp apply (rule_tac x="fromEnum x" in image_eqI) apply clarsimp apply clarsimp apply (rule conjI) apply (subst word_le_nat_alt [symmetric]) apply simp apply safe apply (simp add: word_le_nat_alt [symmetric]) apply (simp add: word_less_nat_alt [symmetric]) done lemma upto_enum_less: assumes xin: "x \ set [(a::'a::len word).e.2 ^ n - 1]" and nv: "n < LENGTH('a::len)" shows "x < 2 ^ n" proof (cases n) case 0 then show ?thesis using xin by simp next case (Suc m) show ?thesis using xin nv by simp qed lemma upto_enum_len_less: "\ n \ length [a, b .e. c]; n \ 0 \ \ a \ c" unfolding upto_enum_step_def by (simp split: if_split_asm) lemma length_upto_enum_step: fixes x :: "'a :: len word" shows "x \ z \ length [x , y .e. z] = (unat ((z - x) div (y - x))) + 1" unfolding upto_enum_step_def by (simp add: upto_enum_red) lemma map_length_unfold_one: fixes x :: "'a::len word" assumes xv: "Suc (unat x) < 2 ^ LENGTH('a)" and ax: "a < x" shows "map f [a .e. x] = f a # map f [a + 1 .e. x]" by (subst word_upto_Cons_eq, auto, fact+) lemma upto_enum_set_conv2: fixes a :: "'a::len word" shows "set [a .e. b] = {a .. b}" by auto lemma of_nat_unat [simp]: "of_nat \ unat = id" by (rule ext, simp) lemma Suc_unat_minus_one [simp]: "x \ 0 \ Suc (unat (x - 1)) = unat x" by (metis Suc_diff_1 unat_gt_0 unat_minus_one) lemma word_add_le_dest: fixes i :: "'a :: len word" assumes le: "i + k \ j + k" and uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i \ j" using uik ujk le by (auto simp: word_le_nat_alt iffD1 [OF unat_add_lem] elim: add_le_mono1) lemma mask_shift: "(x && ~~ mask y) >> y = x >> y" apply (rule word_eqI) apply (simp add: nth_shiftr word_size) apply safe apply (drule test_bit.Rep[simplified, rule_format]) apply (simp add: word_size word_ops_nth_size) done lemma word_add_le_mono1: fixes i :: "'a :: len word" assumes ij: "i \ j" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i + k \ j + k" proof - from ij ujk have jk: "unat i + unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_le_nat_alt elim: add_le_mono1) then show ?thesis using ujk ij by (auto simp: word_le_nat_alt iffD1 [OF unat_add_lem]) qed lemma word_add_le_mono2: fixes i :: "'a :: len word" shows "\i \ j; unat j + unat k < 2 ^ LENGTH('a)\ \ k + i \ k + j" by (subst field_simps, subst field_simps, erule (1) word_add_le_mono1) lemma word_add_le_iff: fixes i :: "'a :: len word" assumes uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "(i + k \ j + k) = (i \ j)" proof assume "i \ j" show "i + k \ j + k" by (rule word_add_le_mono1) fact+ next assume "i + k \ j + k" show "i \ j" by (rule word_add_le_dest) fact+ qed lemma word_add_less_mono1: fixes i :: "'a :: len word" assumes ij: "i < j" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i + k < j + k" proof - from ij ujk have jk: "unat i + unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_less_nat_alt elim: add_less_mono1) then show ?thesis using ujk ij by (auto simp: word_less_nat_alt iffD1 [OF unat_add_lem]) qed lemma word_add_less_dest: fixes i :: "'a :: len word" assumes le: "i + k < j + k" and uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i < j" using uik ujk le by (auto simp: word_less_nat_alt iffD1 [OF unat_add_lem] elim: add_less_mono1) lemma word_add_less_iff: fixes i :: "'a :: len word" assumes uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "(i + k < j + k) = (i < j)" proof assume "i < j" show "i + k < j + k" by (rule word_add_less_mono1) fact+ next assume "i + k < j + k" show "i < j" by (rule word_add_less_dest) fact+ qed lemma shiftr_div_2n': "unat (w >> n) = unat w div 2 ^ n" apply (unfold unat_def) apply (subst shiftr_div_2n) apply (subst nat_div_distrib) apply simp apply (simp add: nat_power_eq) done lemma shiftl_shiftr_id: assumes nv: "n < LENGTH('a)" and xv: "x < 2 ^ (LENGTH('a) - n)" shows "x << n >> n = (x::'a::len word)" apply (simp add: shiftl_t2n) apply (rule word_unat.Rep_eqD) apply (subst shiftr_div_2n') apply (cases n) apply simp apply (subst iffD1 [OF unat_mult_lem])+ apply (subst unat_power_lower[OF nv]) apply (rule nat_less_power_trans [OF _ order_less_imp_le [OF nv]]) apply (rule order_less_le_trans [OF unat_mono [OF xv] order_eq_refl]) apply (rule unat_power_lower) apply simp apply (subst unat_power_lower[OF nv]) apply simp done lemma word_mult_less_iff: fixes i :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k < j * k) = (i < j)" using assms by (rule word_mult_less_cancel) lemma word_le_imp_diff_le: fixes n :: "'a::len word" shows "\k \ n; n \ m\ \ n - k \ m" by (auto simp: unat_sub word_le_nat_alt) lemma word_less_imp_diff_less: fixes n :: "'a::len word" shows "\k \ n; n < m\ \ n - k < m" by (clarsimp simp: unat_sub word_less_nat_alt intro!: less_imp_diff_less) lemma word_mult_le_mono1: fixes i :: "'a :: len word" assumes ij: "i \ j" and knz: "0 < k" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i * k \ j * k" proof - from ij ujk knz have jk: "unat i * unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_le_nat_alt elim: mult_le_mono1) then show ?thesis using ujk knz ij by (auto simp: word_le_nat_alt iffD1 [OF unat_mult_lem]) qed lemma word_mult_le_iff: fixes i :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k \ j * k) = (i \ j)" proof assume "i \ j" show "i * k \ j * k" by (rule word_mult_le_mono1) fact+ next assume p: "i * k \ j * k" have "0 < unat k" using knz by (simp add: word_less_nat_alt) then show "i \ j" using p by (clarsimp simp: word_le_nat_alt iffD1 [OF unat_mult_lem uik] iffD1 [OF unat_mult_lem ujk]) qed lemma word_diff_less: fixes n :: "'a :: len word" shows "\0 < n; 0 < m; n \ m\ \ m - n < m" apply (subst word_less_nat_alt) apply (subst unat_sub) apply assumption apply (rule diff_less) apply (simp_all add: word_less_nat_alt) done lemma MinI: assumes fa: "finite A" and ne: "A \ {}" and xv: "m \ A" and min: "\y \ A. m \ y" shows "Min A = m" using fa ne xv min proof (induct A arbitrary: m rule: finite_ne_induct) case singleton then show ?case by simp next case (insert y F) from insert.prems have yx: "m \ y" and fx: "\y \ F. m \ y" by auto have "m \ insert y F" by fact then show ?case proof assume mv: "m = y" have mlt: "m \ Min F" by (rule iffD2 [OF Min_ge_iff [OF insert.hyps(1) insert.hyps(2)] fx]) show ?case apply (subst Min_insert [OF insert.hyps(1) insert.hyps(2)]) apply (subst mv [symmetric]) apply (auto simp: min_def mlt) done next assume "m \ F" then have mf: "Min F = m" by (rule insert.hyps(4) [OF _ fx]) show ?case apply (subst Min_insert [OF insert.hyps(1) insert.hyps(2)]) apply (subst mf) apply (rule iffD2 [OF _ yx]) apply (auto simp: min_def) done qed qed lemma length_upto_enum [simp]: fixes a :: "'a :: len word" shows "length [a .e. b] = Suc (unat b) - unat a" apply (simp add: word_le_nat_alt upto_enum_red) apply (clarsimp simp: Suc_diff_le) done lemma length_upto_enum_less_one: "\a \ b; b \ 0\ \ length [a .e. b - 1] = unat (b - a)" apply clarsimp apply (subst unat_sub[symmetric], assumption) apply clarsimp done lemma drop_upto_enum: "drop (unat n) [0 .e. m] = [n .e. m]" apply (clarsimp simp: upto_enum_def) apply (induct m, simp) by (metis drop_map drop_upt plus_nat.add_0) lemma distinct_enum_upto' [simp]: "distinct [a::'a::len word .e. b]" apply (subst drop_upto_enum [symmetric]) apply (rule distinct_drop) apply (rule distinct_enum_upto) done lemma length_interval: "\set xs = {x. (a::'a::len word) \ x \ x \ b}; distinct xs\ \ length xs = Suc (unat b) - unat a" apply (frule distinct_card) apply (subgoal_tac "set xs = set [a .e. b]") apply (cut_tac distinct_card [where xs="[a .e. b]"]) apply (subst (asm) length_upto_enum) apply clarsimp apply (rule distinct_enum_upto') apply simp done lemma not_empty_eq: "(S \ {}) = (\x. x \ S)" by auto lemma range_subset_lower: fixes c :: "'a ::linorder" shows "\ {a..b} \ {c..d}; x \ {a..b} \ \ c \ a" apply (frule (1) subsetD) apply (rule classical) apply clarsimp done lemma range_subset_upper: fixes c :: "'a ::linorder" shows "\ {a..b} \ {c..d}; x \ {a..b} \ \ b \ d" apply (frule (1) subsetD) apply (rule classical) apply clarsimp done lemma range_subset_eq: fixes a::"'a::linorder" assumes non_empty: "a \ b" shows "({a..b} \ {c..d}) = (c \ a \ b \ d)" apply (insert non_empty) apply (rule iffI) apply (frule range_subset_lower [where x=a], simp) apply (drule range_subset_upper [where x=a], simp) apply simp apply auto done lemma range_eq: fixes a::"'a::linorder" assumes non_empty: "a \ b" shows "({a..b} = {c..d}) = (a = c \ b = d)" by (metis atLeastatMost_subset_iff eq_iff non_empty) lemma range_strict_subset_eq: fixes a::"'a::linorder" assumes non_empty: "a \ b" shows "({a..b} \ {c..d}) = (c \ a \ b \ d \ (a = c \ b \ d))" apply (insert non_empty) apply (subst psubset_eq) apply (subst range_subset_eq, assumption+) apply (subst range_eq, assumption+) apply simp done lemma range_subsetI: fixes x :: "'a :: order" assumes xX: "X \ x" and yY: "y \ Y" shows "{x .. y} \ {X .. Y}" using xX yY by auto lemma set_False [simp]: "(set bs \ {False}) = (True \ set bs)" by auto declare of_nat_power [simp del] (* TODO: move to word *) lemma unat_of_bl_length: "unat (of_bl xs :: 'a::len word) < 2 ^ (length xs)" proof (cases "length xs < LENGTH('a)") case True then have "(of_bl xs::'a::len word) < 2 ^ length xs" by (simp add: of_bl_length_less) with True show ?thesis by (simp add: word_less_nat_alt word_unat_power unat_of_nat) next case False have "unat (of_bl xs::'a::len word) < 2 ^ LENGTH('a)" by (simp split: unat_split) also from False have "LENGTH('a) \ length xs" by simp then have "2 ^ LENGTH('a) \ (2::nat) ^ length xs" by (rule power_increasing) simp finally show ?thesis . qed lemma is_aligned_0'[simp]: "is_aligned 0 n" by (simp add: is_aligned_def) lemma p_assoc_help: fixes p :: "'a::{ring,power,numeral,one}" shows "p + 2^sz - 1 = p + (2^sz - 1)" by simp lemma word_add_increasing: fixes x :: "'a :: len word" shows "\ p + w \ x; p \ p + w \ \ p \ x" by unat_arith lemma word_random: fixes x :: "'a :: len word" shows "\ p \ p + x'; x \ x' \ \ p \ p + x" by unat_arith lemma word_sub_mono: "\ a \ c; d \ b; a - b \ a; c - d \ c \ \ (a - b) \ (c - d :: 'a :: len word)" by unat_arith lemma power_not_zero: "n < LENGTH('a::len) \ (2 :: 'a word) ^ n \ 0" by (metis p2_gt_0 word_neq_0_conv) lemma word_gt_a_gt_0: "a < n \ (0 :: 'a::len word) < n" apply (case_tac "n = 0") apply clarsimp apply (clarsimp simp: word_neq_0_conv) done lemma word_shift_nonzero: "\ (x::'a::len word) \ 2 ^ m; m + n < LENGTH('a::len); x \ 0\ \ x << n \ 0" apply (simp only: word_neq_0_conv word_less_nat_alt shiftl_t2n mod_0 unat_word_ariths unat_power_lower word_le_nat_alt) apply (subst mod_less) apply (rule order_le_less_trans) apply (erule mult_le_mono2) apply (subst power_add[symmetric]) apply (rule power_strict_increasing) apply simp apply simp apply simp done lemma word_power_less_1 [simp]: "sz < LENGTH('a::len) \ (2::'a word) ^ sz - 1 < 2 ^ sz" apply (simp add: word_less_nat_alt) apply (subst unat_minus_one) apply (simp add: word_unat.Rep_inject [symmetric]) apply simp done lemma nasty_split_lt: "\ (x :: 'a:: len word) < 2 ^ (m - n); n \ m; m < LENGTH('a::len) \ \ x * 2 ^ n + (2 ^ n - 1) \ 2 ^ m - 1" apply (simp only: add_diff_eq) apply (subst mult_1[symmetric], subst distrib_right[symmetric]) apply (rule word_sub_mono) apply (rule order_trans) apply (rule word_mult_le_mono1) apply (rule inc_le) apply assumption apply (subst word_neq_0_conv[symmetric]) apply (rule power_not_zero) apply simp apply (subst unat_power_lower, simp)+ apply (subst power_add[symmetric]) apply (rule power_strict_increasing) apply simp apply simp apply (subst power_add[symmetric]) apply simp apply simp apply (rule word_sub_1_le) apply (subst mult.commute) apply (subst shiftl_t2n[symmetric]) apply (rule word_shift_nonzero) apply (erule inc_le) apply simp apply (unat_arith) apply (drule word_power_less_1) apply simp done lemma nasty_split_less: "\m \ n; n \ nm; nm < LENGTH('a::len); x < 2 ^ (nm - n)\ \ (x :: 'a word) * 2 ^ n + (2 ^ m - 1) < 2 ^ nm" apply (simp only: word_less_sub_le[symmetric]) apply (rule order_trans [OF _ nasty_split_lt]) apply (rule word_plus_mono_right) apply (rule word_sub_mono) apply (simp add: word_le_nat_alt) apply simp apply (simp add: word_sub_1_le[OF power_not_zero]) apply (simp add: word_sub_1_le[OF power_not_zero]) apply (rule is_aligned_no_wrap') apply (rule is_aligned_mult_triv2) apply simp apply (erule order_le_less_trans, simp) apply simp+ done lemma int_not_emptyD: "A \ B \ {} \ \x. x \ A \ x \ B" by (erule contrapos_np, clarsimp simp: disjoint_iff_not_equal) lemma unat_less_power: fixes k :: "'a::len word" assumes szv: "sz < LENGTH('a)" and kv: "k < 2 ^ sz" shows "unat k < 2 ^ sz" using szv unat_mono [OF kv] by simp lemma unat_mult_power_lem: assumes kv: "k < 2 ^ (LENGTH('a::len) - sz)" shows "unat (2 ^ sz * of_nat k :: (('a::len) word)) = 2 ^ sz * k" proof cases assume szv: "sz < LENGTH('a::len)" show ?thesis proof (cases "sz = 0") case True then show ?thesis using kv szv by (simp add: unat_of_nat) next case False then have sne: "0 < sz" .. have uk: "unat (of_nat k :: 'a word) = k" apply (subst unat_of_nat) apply (simp add: nat_mod_eq less_trans[OF kv] sne) done show ?thesis using szv apply (subst iffD1 [OF unat_mult_lem]) apply (simp add: uk nat_less_power_trans[OF kv order_less_imp_le [OF szv]])+ done qed next assume "\ sz < LENGTH('a)" with kv show ?thesis by (simp add: not_less power_overflow) qed lemma aligned_add_offset_no_wrap: fixes off :: "('a::len) word" and x :: "'a word" assumes al: "is_aligned x sz" and offv: "off < 2 ^ sz" shows "unat x + unat off < 2 ^ LENGTH('a)" proof cases assume szv: "sz < LENGTH('a)" from al obtain k where xv: "x = 2 ^ sz * (of_nat k)" and kl: "k < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) show ?thesis using szv apply (subst xv) apply (subst unat_mult_power_lem[OF kl]) apply (subst mult.commute, rule nat_add_offset_less) apply (rule less_le_trans[OF unat_mono[OF offv, simplified]]) apply (erule eq_imp_le[OF unat_power_lower]) apply (rule kl) apply simp done next assume "\ sz < LENGTH('a)" with offv show ?thesis by (simp add: not_less power_overflow ) qed lemma aligned_add_offset_mod: fixes x :: "('a::len) word" assumes al: "is_aligned x sz" and kv: "k < 2 ^ sz" shows "(x + k) mod 2 ^ sz = k" proof cases assume szv: "sz < LENGTH('a)" have ux: "unat x + unat k < 2 ^ LENGTH('a)" by (rule aligned_add_offset_no_wrap) fact+ show ?thesis using al szv apply - apply (erule is_alignedE) apply (subst word_unat.Rep_inject [symmetric]) apply (subst unat_mod) apply (subst iffD1 [OF unat_add_lem], rule ux) apply simp apply (subst unat_mult_power_lem, assumption+) apply (simp) apply (rule mod_less[OF less_le_trans[OF unat_mono], OF kv]) apply (erule eq_imp_le[OF unat_power_lower]) done next assume "\ sz < LENGTH('a)" with al show ?thesis by (simp add: not_less power_overflow is_aligned_mask mask_def word_mod_by_0) qed lemma word_plus_mcs_4: "\v + x \ w + x; x \ v + x\ \ v \ (w::'a::len word)" by uint_arith lemma word_plus_mcs_3: "\v \ w; x \ w + x\ \ v + x \ w + (x::'a::len word)" by unat_arith lemma aligned_neq_into_no_overlap: fixes x :: "'a::len word" assumes neq: "x \ y" and alx: "is_aligned x sz" and aly: "is_aligned y sz" shows "{x .. x + (2 ^ sz - 1)} \ {y .. y + (2 ^ sz - 1)} = {}" proof cases assume szv: "sz < LENGTH('a)" show ?thesis proof (rule equals0I, clarsimp) fix z assume xb: "x \ z" and xt: "z \ x + (2 ^ sz - 1)" and yb: "y \ z" and yt: "z \ y + (2 ^ sz - 1)" have rl: "\(p::'a word) k w. \uint p + uint k < 2 ^ LENGTH('a); w = p + k; w \ p + (2 ^ sz - 1) \ \ k < 2 ^ sz" apply - apply simp apply (subst (asm) add.commute, subst (asm) add.commute, drule word_plus_mcs_4) apply (subst add.commute, subst no_plus_overflow_uint_size) apply (simp add: word_size_bl) apply (erule iffD1 [OF word_less_sub_le[OF szv]]) done from xb obtain kx where kx: "z = x + kx" and kxl: "uint x + uint kx < 2 ^ LENGTH('a)" by (clarsimp dest!: word_le_exists') from yb obtain ky where ky: "z = y + ky" and kyl: "uint y + uint ky < 2 ^ LENGTH('a)" by (clarsimp dest!: word_le_exists') have "x = y" proof - have "kx = z mod 2 ^ sz" proof (subst kx, rule sym, rule aligned_add_offset_mod) show "kx < 2 ^ sz" by (rule rl) fact+ qed fact+ also have "\ = ky" proof (subst ky, rule aligned_add_offset_mod) show "ky < 2 ^ sz" using kyl ky yt by (rule rl) qed fact+ finally have kxky: "kx = ky" . moreover have "x + kx = y + ky" by (simp add: kx [symmetric] ky [symmetric]) ultimately show ?thesis by simp qed then show False using neq by simp qed next assume "\ sz < LENGTH('a)" with neq alx aly have False by (simp add: is_aligned_mask mask_def power_overflow) then show ?thesis .. qed lemma less_two_pow_divD: "\ (x :: nat) < 2 ^ n div 2 ^ m \ \ n \ m \ (x < 2 ^ (n - m))" apply (rule context_conjI) apply (rule ccontr) apply (simp add: power_strict_increasing) apply (simp add: power_sub) done lemma less_two_pow_divI: "\ (x :: nat) < 2 ^ (n - m); m \ n \ \ x < 2 ^ n div 2 ^ m" by (simp add: power_sub) lemma word_less_two_pow_divI: "\ (x :: 'a::len word) < 2 ^ (n - m); m \ n; n < LENGTH('a) \ \ x < 2 ^ n div 2 ^ m" apply (simp add: word_less_nat_alt) apply (subst unat_word_ariths) apply (subst mod_less) apply (rule order_le_less_trans [OF div_le_dividend]) apply (rule unat_lt2p) apply (simp add: power_sub) done lemma word_less_two_pow_divD: "\ (x :: 'a::len word) < 2 ^ n div 2 ^ m \ \ n \ m \ (x < 2 ^ (n - m))" apply (cases "n < LENGTH('a)") apply (cases "m < LENGTH('a)") apply (simp add: word_less_nat_alt) apply (subst(asm) unat_word_ariths) apply (subst(asm) mod_less) apply (rule order_le_less_trans [OF div_le_dividend]) apply (rule unat_lt2p) apply (clarsimp dest!: less_two_pow_divD) apply (simp add: power_overflow) apply (simp add: word_div_def) apply (simp add: power_overflow word_div_def) done lemma of_nat_less_two_pow_div_set: "\ n < LENGTH('a) \ \ {x. x < (2 ^ n div 2 ^ m :: 'a::len word)} = of_nat ` {k. k < 2 ^ n div 2 ^ m}" apply (simp add: image_def) apply (safe dest!: word_less_two_pow_divD less_two_pow_divD intro!: word_less_two_pow_divI) apply (rule_tac x="unat x" in exI) apply (simp add: power_sub[symmetric]) apply (subst unat_power_lower[symmetric, where 'a='a]) apply simp apply (erule unat_mono) apply (subst word_unat_power) apply (rule of_nat_mono_maybe) apply (rule power_strict_increasing) apply simp apply simp apply assumption done lemma word_less_power_trans2: fixes n :: "'a::len word" shows "\n < 2 ^ (m - k); k \ m; m < LENGTH('a)\ \ n * 2 ^ k < 2 ^ m" by (subst field_simps, rule word_less_power_trans) lemma ucast_less: "LENGTH('b) < LENGTH('a) \ (ucast (x :: 'b :: len word) :: ('a :: len word)) < 2 ^ LENGTH('b)" apply (subst mask_eq_iff_w2p[symmetric]) apply (simp add: word_size) apply (rule word_eqI) apply (simp add: word_size nth_ucast) apply safe apply (simp add: test_bit.Rep[simplified]) done lemma ucast_range_less: "LENGTH('a :: len) < LENGTH('b :: len) \ range (ucast :: 'a word \ 'b word) = {x. x < 2 ^ len_of TYPE ('a)}" apply safe apply (erule ucast_less) apply (simp add: image_def) apply (rule_tac x="ucast x" in exI) apply (drule less_mask_eq) apply (rule word_eqI) apply (drule_tac x=n in word_eqD) apply (clarsimp simp: word_size nth_ucast) done lemma word_power_less_diff: "\2 ^ n * q < (2::'a::len word) ^ m; q < 2 ^ (LENGTH('a) - n)\ \ q < 2 ^ (m - n)" apply (case_tac "m \ LENGTH('a)") apply (simp add: power_overflow) apply (case_tac "n \ LENGTH('a)") apply (simp add: power_overflow) apply (cases "n = 0") apply simp apply (subst word_less_nat_alt) apply (subst unat_power_lower) apply simp apply (rule nat_power_less_diff) apply (simp add: word_less_nat_alt) apply (subst (asm) iffD1 [OF unat_mult_lem]) apply (simp add:nat_less_power_trans) apply simp done lemmas word_diff_ls'' = word_diff_ls [where xa=x and x=x for x] lemmas word_diff_ls' = word_diff_ls'' [simplified] lemmas word_l_diffs' = word_l_diffs [where xa=x and x=x for x] lemmas word_l_diffs = word_l_diffs' [simplified] lemma is_aligned_diff: fixes m :: "'a::len word" assumes alm: "is_aligned m s1" and aln: "is_aligned n s2" and s2wb: "s2 < LENGTH('a)" and nm: "m \ {n .. n + (2 ^ s2 - 1)}" and s1s2: "s1 \ s2" and s10: "0 < s1" (* Probably can be folded into the proof \ *) shows "\q. m - n = of_nat q * 2 ^ s1 \ q < 2 ^ (s2 - s1)" proof - have rl: "\m s. \ m < 2 ^ (LENGTH('a) - s); s < LENGTH('a) \ \ unat ((2::'a word) ^ s * of_nat m) = 2 ^ s * m" proof - fix m :: nat and s assume m: "m < 2 ^ (LENGTH('a) - s)" and s: "s < LENGTH('a)" then have "unat ((of_nat m) :: 'a word) = m" apply (subst unat_of_nat) apply (subst mod_less) apply (erule order_less_le_trans) apply (rule power_increasing) apply simp_all done then show "?thesis m s" using s m apply (subst iffD1 [OF unat_mult_lem]) apply (simp add: nat_less_power_trans)+ done qed have s1wb: "s1 < LENGTH('a)" using s2wb s1s2 by simp from alm obtain mq where mmq: "m = 2 ^ s1 * of_nat mq" and mq: "mq < 2 ^ (LENGTH('a) - s1)" by (auto elim: is_alignedE simp: field_simps) from aln obtain nq where nnq: "n = 2 ^ s2 * of_nat nq" and nq: "nq < 2 ^ (LENGTH('a) - s2)" by (auto elim: is_alignedE simp: field_simps) from s1s2 obtain sq where sq: "s2 = s1 + sq" by (auto simp: le_iff_add) note us1 = rl [OF mq s1wb] note us2 = rl [OF nq s2wb] from nm have "n \ m" by clarsimp then have "(2::'a word) ^ s2 * of_nat nq \ 2 ^ s1 * of_nat mq" using nnq mmq by simp then have "2 ^ s2 * nq \ 2 ^ s1 * mq" using s1wb s2wb by (simp add: word_le_nat_alt us1 us2) then have nqmq: "2 ^ sq * nq \ mq" using sq by (simp add: power_add) have "m - n = 2 ^ s1 * of_nat mq - 2 ^ s2 * of_nat nq" using mmq nnq by simp also have "\ = 2 ^ s1 * of_nat mq - 2 ^ s1 * 2 ^ sq * of_nat nq" using sq by (simp add: power_add) also have "\ = 2 ^ s1 * (of_nat mq - 2 ^ sq * of_nat nq)" by (simp add: field_simps) also have "\ = 2 ^ s1 * of_nat (mq - 2 ^ sq * nq)" using s1wb s2wb us1 us2 nqmq by (simp add: word_unat_power) finally have mn: "m - n = of_nat (mq - 2 ^ sq * nq) * 2 ^ s1" by simp moreover from nm have "m - n \ 2 ^ s2 - 1" by - (rule word_diff_ls', (simp add: field_simps)+) then have "(2::'a word) ^ s1 * of_nat (mq - 2 ^ sq * nq) < 2 ^ s2" using mn s2wb by (simp add: field_simps) then have "of_nat (mq - 2 ^ sq * nq) < (2::'a word) ^ (s2 - s1)" proof (rule word_power_less_diff) have mm: "mq - 2 ^ sq * nq < 2 ^ (LENGTH('a) - s1)" using mq by simp moreover from s10 have "LENGTH('a) - s1 < LENGTH('a)" by (rule diff_less, simp) ultimately show "of_nat (mq - 2 ^ sq * nq) < (2::'a word) ^ (LENGTH('a) - s1)" apply (simp add: word_less_nat_alt) apply (subst unat_of_nat) apply (subst mod_less) apply (erule order_less_le_trans) apply simp+ done qed then have "mq - 2 ^ sq * nq < 2 ^ (s2 - s1)" using mq s2wb apply (simp add: word_less_nat_alt) apply (subst (asm) unat_of_nat) apply (subst (asm) mod_less) apply (rule order_le_less_trans) apply (rule diff_le_self) apply (erule order_less_le_trans) apply simp apply assumption done ultimately show ?thesis by auto qed lemma word_less_sub_1: "x < (y :: 'a :: len word) \ x \ y - 1" apply (erule udvd_minus_le') apply (simp add: udvd_def)+ done lemma word_sub_mono2: "\ a + b \ c + d; c \ a; b \ a + b; d \ c + d \ \ b \ (d :: 'a :: len word)" apply (drule(1) word_sub_mono) apply simp apply simp apply simp done lemma word_not_le: "(\ x \ (y :: 'a :: len word)) = (y < x)" by fastforce lemma word_subset_less: "\ {x .. x + r - 1} \ {y .. y + s - 1}; x \ x + r - 1; y \ y + (s :: 'a :: len word) - 1; s \ 0 \ \ r \ s" apply (frule subsetD[where c=x]) apply simp apply (drule subsetD[where c="x + r - 1"]) apply simp apply (clarsimp simp: add_diff_eq[symmetric]) apply (drule(1) word_sub_mono2) apply (simp_all add: olen_add_eqv[symmetric]) apply (erule word_le_minus_cancel) apply (rule ccontr) apply (simp add: word_not_le) done lemma uint_power_lower: "n < LENGTH('a) \ uint (2 ^ n :: 'a :: len word) = (2 ^ n :: int)" by (rule uint_2p_alt) lemma power_le_mono: "\2 ^ n \ (2::'a::len word) ^ m; n < LENGTH('a); m < LENGTH('a)\ \ n \ m" apply (clarsimp simp add: le_less) apply safe apply (simp add: word_less_nat_alt) apply (simp only: uint_arith_simps(3)) apply (drule uint_power_lower)+ apply simp done lemma sublist_equal_part: "prefix xs ys \ take (length xs) ys = xs" by (clarsimp simp: prefix_def) lemma two_power_eq: "\n < LENGTH('a); m < LENGTH('a)\ \ ((2::'a::len word) ^ n = 2 ^ m) = (n = m)" apply safe apply (rule order_antisym) apply (simp add: power_le_mono[where 'a='a])+ done lemma prefix_length_less: "strict_prefix xs ys \ length xs < length ys" apply (clarsimp simp: strict_prefix_def) apply (frule prefix_length_le) apply (rule ccontr, simp) apply (clarsimp simp: prefix_def) done lemmas take_less = take_strict_prefix lemma not_prefix_longer: "\ length xs > length ys \ \ \ prefix xs ys" by (clarsimp dest!: prefix_length_le) lemma of_bl_length: "length xs < LENGTH('a) \ of_bl xs < (2 :: 'a::len word) ^ length xs" by (simp add: of_bl_length_less) lemma unat_of_nat_eq: "x < 2 ^ LENGTH('a) \ unat (of_nat x ::'a::len word) = x" by (rule unat_of_nat_len) lemma unat_eq_of_nat: "n < 2 ^ LENGTH('a) \ (unat (x :: 'a::len word) = n) = (x = of_nat n)" by (subst unat_of_nat_eq[where x=n, symmetric], simp+) lemma unat_less_helper: "x < of_nat n \ unat x < n" apply (simp add: word_less_nat_alt) apply (erule order_less_le_trans) apply (simp add: unat_of_nat) done lemma of_nat_0: "\of_nat n = (0::('a::len) word); n < 2 ^ len_of (TYPE('a))\ \ n = 0" by (drule unat_of_nat_eq, simp) lemma minus_one_helper3: "x < y \ x \ (y :: 'a :: len word) - 1" apply (simp add: word_less_nat_alt word_le_nat_alt) apply (subst unat_minus_one) apply clarsimp apply arith done lemma minus_one_helper: "\ x \ y; x \ 0 \ \ x - 1 < (y :: 'a :: len word)" apply (simp add: word_less_nat_alt word_le_nat_alt) apply (subst unat_minus_one) apply assumption apply (cases "unat x") apply (simp add: unat_eq_zero) apply arith done lemma minus_one_helper5: fixes x :: "'a::len word" shows "\y \ 0; x \ y - 1 \ \ x < y" using le_m1_iff_lt word_neq_0_conv by blast lemma not_greatest_aligned: "\ x < y; is_aligned x n; is_aligned y n \ \ x + 2 ^ n \ 0" apply (rule notI) apply (erule is_aligned_get_word_bits[where p=y]) apply (simp add: eq_diff_eq[symmetric]) apply (frule minus_one_helper3) apply (drule le_minus'[where a="x" and c="y - x" and b="- 1" for x y, simplified]) apply (simp add: field_simps) apply (frule is_aligned_less_sz[where a=y]) apply clarsimp apply (erule notE) apply (rule minus_one_helper5) apply simp apply (metis is_aligned_no_overflow minus_one_helper3 order_le_less_trans) apply simp done lemma of_nat_inj: "\x < 2 ^ LENGTH('a); y < 2 ^ LENGTH('a)\ \ (of_nat x = (of_nat y :: 'a :: len word)) = (x = y)" by (simp add: word_unat.norm_eq_iff [symmetric]) lemma map_prefixI: "prefix xs ys \ prefix (map f xs) (map f ys)" by (clarsimp simp: prefix_def) lemma if_Some_None_eq_None: "((if P then Some v else None) = None) = (\ P)" by simp lemma CollectPairFalse [iff]: "{(a,b). False} = {}" by (simp add: split_def) lemma if_conj_dist: "((if b then w else x) \ (if b then y else z) \ X) = ((if b then w \ y else x \ z) \ X)" by simp lemma if_P_True1: "Q \ (if P then True else Q)" by simp lemma if_P_True2: "Q \ (if P then Q else True)" by simp lemma list_all2_induct [consumes 1, case_names Nil Cons]: assumes lall: "list_all2 Q xs ys" and nilr: "P [] []" and consr: "\x xs y ys. \list_all2 Q xs ys; Q x y; P xs ys\ \ P (x # xs) (y # ys)" shows "P xs ys" using lall proof (induct rule: list_induct2 [OF list_all2_lengthD [OF lall]]) case 1 then show ?case by auto fact+ next case (2 x xs y ys) show ?case proof (rule consr) from "2.prems" show "list_all2 Q xs ys" and "Q x y" by simp_all then show "P xs ys" by (intro "2.hyps") qed qed lemma list_all2_induct_suffixeq [consumes 1, case_names Nil Cons]: assumes lall: "list_all2 Q as bs" and nilr: "P [] []" and consr: "\x xs y ys. \list_all2 Q xs ys; Q x y; P xs ys; suffix (x # xs) as; suffix (y # ys) bs\ \ P (x # xs) (y # ys)" shows "P as bs" proof - define as' where "as' == as" define bs' where "bs' == bs" have "suffix as as' \ suffix bs bs'" unfolding as'_def bs'_def by simp then show ?thesis using lall proof (induct rule: list_induct2 [OF list_all2_lengthD [OF lall]]) case 1 show ?case by fact next case (2 x xs y ys) show ?case proof (rule consr) from "2.prems" show "list_all2 Q xs ys" and "Q x y" by simp_all then show "P xs ys" using "2.hyps" "2.prems" by (auto dest: suffix_ConsD) from "2.prems" show "suffix (x # xs) as" and "suffix (y # ys) bs" by (auto simp: as'_def bs'_def) qed qed qed lemma upto_enum_step_shift: "\ is_aligned p n \ \ ([p , p + 2 ^ m .e. p + 2 ^ n - 1]) = map ((+) p) [0, 2 ^ m .e. 2 ^ n - 1]" apply (erule is_aligned_get_word_bits) prefer 2 apply (simp add: map_idI) apply (clarsimp simp: upto_enum_step_def) apply (frule is_aligned_no_overflow) apply (simp add: linorder_not_le [symmetric]) done lemma upto_enum_step_shift_red: "\ is_aligned p sz; sz < len_of (TYPE('a)); us \ sz \ \ [p :: 'a :: len word, p + 2 ^ us .e. p + 2 ^ sz - 1] = map (\x. p + of_nat x * 2 ^ us) [0 ..< 2 ^ (sz - us)]" apply (subst upto_enum_step_shift, assumption) apply (simp add: upto_enum_step_red) done lemma div_to_mult_word_lt: "\ (x :: 'a :: len word) \ y div z \ \ x * z \ y" apply (cases "z = 0") apply simp apply (simp add: word_neq_0_conv) apply (rule order_trans) apply (erule(1) word_mult_le_mono1) apply (simp add: unat_div) apply (rule order_le_less_trans [OF div_mult_le]) apply simp apply (rule word_div_mult_le) done lemma upto_enum_step_subset: "set [x, y .e. z] \ {x .. z}" apply (clarsimp simp: upto_enum_step_def linorder_not_less) apply (drule div_to_mult_word_lt) apply (rule conjI) apply (erule word_random[rotated]) apply simp apply (rule order_trans) apply (erule word_plus_mono_right) apply simp apply simp done lemma shiftr_less_t2n': fixes x :: "'a :: len word" shows "\ x && mask (n + m) = x; m < LENGTH('a) \ \ (x >> n) < 2 ^ m" apply (subst mask_eq_iff_w2p[symmetric]) apply (simp add: word_size) apply (rule word_eqI) apply (drule_tac x="na + n" in word_eqD) apply (simp add: nth_shiftr word_size) apply safe done lemma shiftr_less_t2n: fixes x :: "'a :: len word" shows "x < 2 ^ (n + m) \ (x >> n) < 2 ^ m" apply (rule shiftr_less_t2n') apply (erule less_mask_eq) apply (rule ccontr) apply (simp add: not_less) apply (subst (asm) p2_eq_0[symmetric]) apply (simp add: power_add) done lemma shiftr_eq_0: "n \ LENGTH('a) \ ((w::'a::len word) >> n) = 0" apply (cut_tac shiftr_less_t2n'[of w n 0], simp) apply (simp add: mask_eq_iff) apply (simp add: lt2p_lem) apply simp done lemma shiftr_not_mask_0: "n+m \ LENGTH('a :: len) \ ((w::'a::len word) >> n) && ~~ mask m = 0" apply (simp add: and_not_mask shiftr_less_t2n shiftr_shiftr) apply (subgoal_tac "w >> n + m = 0", simp) apply (simp add: le_mask_iff[symmetric] mask_def le_def) apply (subst (asm) p2_gt_0[symmetric]) apply (simp add: power_add not_less) done lemma shiftl_less_t2n: fixes x :: "'a :: len word" shows "\ x < (2 ^ (m - n)); m < LENGTH('a) \ \ (x << n) < 2 ^ m" apply (subst mask_eq_iff_w2p[symmetric]) apply (simp add: word_size) apply (drule less_mask_eq) apply (rule word_eqI) apply (drule_tac x="na - n" in word_eqD) apply (simp add: nth_shiftl word_size) apply (cases "n \ m") apply safe apply simp apply simp done lemma shiftl_less_t2n': "(x::'a::len word) < 2 ^ m \ m+n < LENGTH('a) \ x << n < 2 ^ (m + n)" by (rule shiftl_less_t2n) simp_all lemma ucast_ucast_mask: "(ucast :: 'a :: len word \ 'b :: len word) (ucast x) = x && mask (len_of TYPE ('a))" apply (rule word_eqI) apply (simp add: nth_ucast word_size) done lemma ucast_ucast_len: "\ x < 2 ^ LENGTH('b) \ \ ucast (ucast x::'b::len word) = (x::'a::len word)" apply (subst ucast_ucast_mask) apply (erule less_mask_eq) done lemma unat_ucast: "unat (ucast x :: ('a :: len0) word) = unat x mod 2 ^ (LENGTH('a))" apply (simp add: unat_def ucast_def) apply (subst word_uint.eq_norm) apply (subst nat_mod_distrib) apply simp apply simp apply (subst nat_power_eq) apply simp apply simp done lemma ucast_less_ucast: "LENGTH('a) < LENGTH('b) \ (ucast x < ((ucast (y :: 'a::len word)) :: 'b::len word)) = (x < y)" apply (simp add: word_less_nat_alt unat_ucast) apply (subst mod_less) apply(rule less_le_trans[OF unat_lt2p], simp) apply (subst mod_less) apply(rule less_le_trans[OF unat_lt2p], simp) apply simp done lemma sints_subset: "m \ n \ sints m \ sints n" apply (simp add: sints_num) apply clarsimp apply (rule conjI) apply (erule order_trans[rotated]) apply simp apply (erule order_less_le_trans) apply simp done lemma up_scast_inj: "\ scast x = (scast y :: 'b :: len word); size x \ LENGTH('b) \ \ x = y" apply (simp add: scast_def) apply (subst(asm) word_sint.Abs_inject) apply (erule subsetD [OF sints_subset]) apply (simp add: word_size) apply (erule subsetD [OF sints_subset]) apply (simp add: word_size) apply simp done lemma up_scast_inj_eq: "LENGTH('a) \ len_of TYPE ('b) \ (scast x = (scast y::'b::len word)) = (x = (y::'a::len word))" by (fastforce dest: up_scast_inj simp: word_size) lemma nth_bounded: "\(x :: 'a :: len word) !! n; x < 2 ^ m; m \ len_of TYPE ('a)\ \ n < m" apply (frule test_bit_size) apply (clarsimp simp: test_bit_bl word_size) apply (simp add: nth_rev) apply (subst(asm) is_aligned_add_conv[OF is_aligned_0', simplified add_0_left, rotated]) apply assumption+ apply (simp only: to_bl_0) apply (simp add: nth_append split: if_split_asm) done lemma is_aligned_add_or: "\is_aligned p n; d < 2 ^ n\ \ p + d = p || d" apply (rule word_plus_and_or_coroll) apply (erule is_aligned_get_word_bits) apply (rule word_eqI) apply (clarsimp simp add: is_aligned_nth) apply (frule(1) nth_bounded) apply simp+ done lemma two_power_increasing: "\ n \ m; m < LENGTH('a) \ \ (2 :: 'a :: len word) ^ n \ 2 ^ m" by (simp add: word_le_nat_alt) lemma is_aligned_add_less_t2n: "\is_aligned (p::'a::len word) n; d < 2^n; n \ m; p < 2^m\ \ p + d < 2^m" apply (case_tac "m < LENGTH('a)") apply (subst mask_eq_iff_w2p[symmetric]) apply (simp add: word_size) apply (simp add: is_aligned_add_or word_ao_dist less_mask_eq) apply (subst less_mask_eq) apply (erule order_less_le_trans) apply (erule(1) two_power_increasing) apply simp apply (simp add: power_overflow) done lemma aligned_offset_non_zero: "\ is_aligned x n; y < 2 ^ n; x \ 0 \ \ x + y \ 0" apply (cases "y = 0") apply simp apply (subst word_neq_0_conv) apply (subst gt0_iff_gem1) apply (erule is_aligned_get_word_bits) apply (subst field_simps[symmetric], subst plus_le_left_cancel_nowrap) apply (rule is_aligned_no_wrap') apply simp apply (rule minus_one_helper) apply simp apply assumption apply (erule (1) is_aligned_no_wrap') apply (simp add: gt0_iff_gem1 [symmetric] word_neq_0_conv) apply simp done lemmas mask_inner_mask = mask_eqs(1) lemma mask_add_aligned: "is_aligned p n \ (p + q) && mask n = q && mask n" apply (simp add: is_aligned_mask) apply (subst mask_inner_mask [symmetric]) apply simp done lemma take_prefix: "(take (length xs) ys = xs) = prefix xs ys" proof (induct xs arbitrary: ys) case Nil then show ?case by simp next case Cons then show ?case by (cases ys) auto qed lemma cart_singleton_empty: "(S \ {e} = {}) = (S = {})" by blast lemma word_div_1: "(n :: 'a :: len word) div 1 = n" by (simp add: word_div_def) lemma word_minus_one_le: "-1 \ (x :: 'a :: len word) = (x = -1)" apply (insert word_n1_ge[where y=x]) apply safe apply (erule(1) order_antisym) done lemma mask_out_sub_mask: "(x && ~~ mask n) = x - (x && mask n)" by (simp add: field_simps word_plus_and_or_coroll2) lemma is_aligned_addD1: assumes al1: "is_aligned (x + y) n" and al2: "is_aligned (x::'a::len word) n" shows "is_aligned y n" using al2 proof (rule is_aligned_get_word_bits) assume "x = 0" then show ?thesis using al1 by simp next assume nv: "n < LENGTH('a)" from al1 obtain q1 where xy: "x + y = 2 ^ n * of_nat q1" and "q1 < 2 ^ (LENGTH('a) - n)" by (rule is_alignedE) moreover from al2 obtain q2 where x: "x = 2 ^ n * of_nat q2" and "q2 < 2 ^ (LENGTH('a) - n)" by (rule is_alignedE) ultimately have "y = 2 ^ n * (of_nat q1 - of_nat q2)" by (simp add: field_simps) then show ?thesis using nv by (simp add: is_aligned_mult_triv1) qed lemmas is_aligned_addD2 = is_aligned_addD1[OF subst[OF add.commute, of "%x. is_aligned x n" for n]] lemma is_aligned_add: "\is_aligned p n; is_aligned q n\ \ is_aligned (p + q) n" by (simp add: is_aligned_mask mask_add_aligned) lemma word_le_add: fixes x :: "'a :: len word" shows "x \ y \ \n. y = x + of_nat n" by (rule exI [where x = "unat (y - x)"]) simp lemma word_plus_mcs_4': fixes x :: "'a :: len word" shows "\x + v \ x + w; x \ x + v\ \ v \ w" apply (rule word_plus_mcs_4) apply (simp add: add.commute) apply (simp add: add.commute) done lemma shiftl_mask_is_0[simp]: "(x << n) && mask n = 0" apply (rule iffD1 [OF is_aligned_mask]) apply (rule is_aligned_shiftl_self) done definition sum_map :: "('a \ 'b) \ ('c \ 'd) \ 'a + 'c \ 'b + 'd" where "sum_map f g x \ case x of Inl v \ Inl (f v) | Inr v' \ Inr (g v')" lemma sum_map_simps[simp]: "sum_map f g (Inl v) = Inl (f v)" "sum_map f g (Inr w) = Inr (g w)" by (simp add: sum_map_def)+ lemma if_and_helper: "(If x v v') && v'' = If x (v && v'') (v' && v'')" by (rule if_distrib) lemma unat_Suc2: fixes n :: "'a :: len word" shows "n \ -1 \ unat (n + 1) = Suc (unat n)" apply (subst add.commute, rule unatSuc) apply (subst eq_diff_eq[symmetric], simp add: minus_equation_iff) done lemmas word_unat_Rep_inject1 = word_unat.Rep_inject[where y=1] lemmas unat_eq_1 = unat_eq_0 word_unat_Rep_inject1[simplified] lemma rshift_sub_mask_eq: "(a >> (size a - b)) && mask b = a >> (size a - b)" using shiftl_shiftr2[where a=a and b=0 and c="size a - b"] apply (cases "b < size a") apply simp apply (simp add: linorder_not_less mask_def word_size p2_eq_0[THEN iffD2]) done lemma shiftl_shiftr3: "b \ c \ a << b >> c = (a >> c - b) && mask (size a - c)" apply (cases "b = c") apply (simp add: shiftl_shiftr1) apply (simp add: shiftl_shiftr2) done lemma and_mask_shiftr_comm: "m\size w \ (w && mask m) >> n = (w >> n) && mask (m-n)" by (simp add: and_mask shiftr_shiftr) (simp add: word_size shiftl_shiftr3) lemma and_mask_shiftl_comm: "m+n \ size w \ (w && mask m) << n = (w << n) && mask (m+n)" by (simp add: and_mask word_size shiftl_shiftl) (simp add: shiftl_shiftr1) lemma le_mask_shiftl_le_mask: "s = m + n \ x \ mask n \ x << m \ mask s" by (simp add: le_mask_iff shiftl_shiftr3) lemma and_not_mask_twice: "(w && ~~ mask n) && ~~ mask m = w && ~~ mask (max m n)" apply (simp add: and_not_mask) apply (case_tac "n x = y - 1 \ x < y - (1 ::'a::len word)" apply (drule word_less_sub_1) apply (drule order_le_imp_less_or_eq) apply auto done lemma eq_eqI: "a = b \ (a = x) = (b = x)" by simp lemma mask_and_mask: "mask a && mask b = mask (min a b)" by (rule word_eqI) (simp add: word_size) lemma mask_eq_0_eq_x: "(x && w = 0) = (x && ~~ w = x)" using word_plus_and_or_coroll2[where x=x and w=w] by auto lemma mask_eq_x_eq_0: "(x && w = x) = (x && ~~ w = 0)" using word_plus_and_or_coroll2[where x=x and w=w] by auto definition "limited_and (x :: 'a :: len word) y = (x && y = x)" lemma limited_and_eq_0: "\ limited_and x z; y && ~~ z = y \ \ x && y = 0" unfolding limited_and_def apply (subst arg_cong2[where f="(&&)"]) apply (erule sym)+ apply (simp(no_asm) add: word_bw_assocs word_bw_comms word_bw_lcs) done lemma limited_and_eq_id: "\ limited_and x z; y && z = z \ \ x && y = x" unfolding limited_and_def by (erule subst, fastforce simp: word_bw_lcs word_bw_assocs word_bw_comms) lemma lshift_limited_and: "limited_and x z \ limited_and (x << n) (z << n)" unfolding limited_and_def by (simp add: shiftl_over_and_dist[symmetric]) lemma rshift_limited_and: "limited_and x z \ limited_and (x >> n) (z >> n)" unfolding limited_and_def by (simp add: shiftr_over_and_dist[symmetric]) lemmas limited_and_simps1 = limited_and_eq_0 limited_and_eq_id lemmas is_aligned_limited_and = is_aligned_neg_mask_eq[unfolded mask_def, folded limited_and_def] lemma compl_of_1: "~~ 1 = (-2 :: 'a :: len word)" by simp lemmas limited_and_simps = limited_and_simps1 limited_and_simps1[OF is_aligned_limited_and] limited_and_simps1[OF lshift_limited_and] limited_and_simps1[OF rshift_limited_and] limited_and_simps1[OF rshift_limited_and, OF is_aligned_limited_and] compl_of_1 shiftl_shiftr1[unfolded word_size mask_def] shiftl_shiftr2[unfolded word_size mask_def] lemma split_word_eq_on_mask: "(x = y) = (x && m = y && m \ x && ~~ m = y && ~~ m)" apply safe apply (rule word_eqI) apply (drule_tac x=n in word_eqD)+ by (auto simp: word_size word_ops_nth_size) lemma map2_Cons_2_3: "(map2 f xs (y # ys) = (z # zs)) = (\x xs'. xs = x # xs' \ f x y = z \ map2 f xs' ys = zs)" by (case_tac xs, simp_all) lemma map2_xor_replicate_False: "map2 (\x y. x \ \ y) xs (replicate n False) = take n xs" apply (induct xs arbitrary: n) apply simp apply (case_tac n) apply simp_all done lemma word_or_zero: "(a || b = 0) = (a = 0 \ b = 0)" apply (safe, simp_all) apply (rule word_eqI, drule_tac x=n in word_eqD, simp)+ done lemma word_and_1_shiftl: fixes x :: "'a :: len word" shows "x && (1 << n) = (if x !! n then (1 << n) else 0)" apply (rule word_eqI) apply (simp add: word_size nth_shiftl del: shiftl_1) apply auto done lemmas word_and_1_shiftls' = word_and_1_shiftl[where n=0] word_and_1_shiftl[where n=1] word_and_1_shiftl[where n=2] lemmas word_and_1_shiftls = word_and_1_shiftls' [simplified] lemma word_and_mask_shiftl: "x && (mask n << m) = ((x >> m) && mask n) << m" apply (rule word_eqI) apply (simp add: word_size nth_shiftl nth_shiftr) apply auto done lemma plus_Collect_helper: "(+) x ` {xa. P (xa :: 'a :: len word)} = {xa. P (xa - x)}" by (fastforce simp add: image_def) lemma plus_Collect_helper2: "(+) (- x) ` {xa. P (xa :: 'a :: len word)} = {xa. P (x + xa)}" using plus_Collect_helper [of "- x" P] by (simp add: ac_simps) lemma word_FF_is_mask: "0xFF = mask 8" by (simp add: mask_def) lemma word_1FF_is_mask: "0x1FF = mask 9" by (simp add: mask_def) lemma ucast_of_nat_small: "x < 2 ^ LENGTH('a) \ ucast (of_nat x :: 'a :: len word) = (of_nat x :: 'b :: len word)" apply (rule sym, subst word_unat.inverse_norm) apply (simp add: ucast_def word_of_int[symmetric] of_nat_nat[symmetric] unat_def[symmetric]) apply (simp add: unat_of_nat) done lemma word_le_make_less: fixes x :: "'a :: len word" shows "y \ -1 \ (x \ y) = (x < (y + 1))" apply safe apply (erule plus_one_helper2) apply (simp add: eq_diff_eq[symmetric]) done lemmas finite_word = finite [where 'a="'a::len word"] lemma word_to_1_set: "{0 ..< (1 :: 'a :: len word)} = {0}" by fastforce lemma range_subset_eq2: "{a :: 'a :: len word .. b} \ {} \ ({a .. b} \ {c .. d}) = (c \ a \ b \ d)" by simp lemma word_count_from_top: "n \ 0 \ {0 ..< n :: 'a :: len word} = {0 ..< n - 1} \ {n - 1}" apply (rule set_eqI, rule iffI) apply simp apply (drule minus_one_helper3) apply (rule disjCI) apply simp apply simp apply (erule minus_one_helper5) apply fastforce done lemma minus_one_helper2: "\ x - 1 < y \ \ x \ (y :: 'a :: len word)" apply (cases "x = 0") apply simp apply (simp add: word_less_nat_alt word_le_nat_alt) apply (subst(asm) unat_minus_one) apply (simp add: word_less_nat_alt) apply (cases "unat x") apply (simp add: unat_eq_zero) apply arith done lemma mod_mod_power: fixes k :: nat shows "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ (min m n)" proof (cases "m \ n") case True then have "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ m" apply - apply (subst mod_less [where n = "2 ^ n"]) apply (rule order_less_le_trans [OF mod_less_divisor]) apply simp+ done also have "\ = k mod 2 ^ (min m n)" using True by simp finally show ?thesis . next case False then have "n < m" by simp then obtain d where md: "m = n + d" by (auto dest: less_imp_add_positive) then have "k mod 2 ^ m = 2 ^ n * (k div 2 ^ n mod 2 ^ d) + k mod 2 ^ n" by (simp add: mod_mult2_eq power_add) then have "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ n" by (simp add: mod_add_left_eq) then show ?thesis using False by simp qed lemma word_div_less: fixes m :: "'a :: len word" shows "m < n \ m div n = 0" apply (rule word_unat.Rep_eqD) apply (simp add: word_less_nat_alt unat_div) done lemma word_must_wrap: "\ x \ n - 1; n \ x \ \ n = (0 :: 'a :: len word)" apply (rule ccontr) apply (drule(1) order_trans) apply (drule word_sub_1_le) apply (drule(1) order_antisym) apply simp done lemma range_subset_card: "\ {a :: 'a :: len word .. b} \ {c .. d}; b \ a \ \ d \ c \ d - c \ b - a" apply (subgoal_tac "a \ {a .. b}") apply (frule(1) range_subset_lower) apply (frule(1) range_subset_upper) apply (rule context_conjI, simp) apply (rule word_sub_mono, assumption+) apply (erule word_sub_le) apply (erule word_sub_le) apply simp done lemma less_1_simp: "n - 1 < m = (n \ (m :: 'a :: len word) \ n \ 0)" by unat_arith lemma alignUp_div_helper: fixes a :: "'a::len word" assumes kv: "k < 2 ^ (LENGTH('a) - n)" and xk: "x = 2 ^ n * of_nat k" and le: "a \ x" and sz: "n < LENGTH('a)" and anz: "a mod 2 ^ n \ 0" shows "a div 2 ^ n < of_nat k" proof - have kn: "unat (of_nat k :: 'a word) * unat ((2::'a word) ^ n) < 2 ^ LENGTH('a)" using xk kv sz apply (subst unat_of_nat_eq) apply (erule order_less_le_trans) apply simp apply (subst unat_power_lower, simp) apply (subst mult.commute) apply (rule nat_less_power_trans) apply simp apply simp done have "unat a div 2 ^ n * 2 ^ n \ unat a" proof - have "unat a = unat a div 2 ^ n * 2 ^ n + unat a mod 2 ^ n" by (simp add: div_mult_mod_eq) also have "\ \ unat a div 2 ^ n * 2 ^ n" using sz anz by (simp add: unat_arith_simps) finally show ?thesis .. qed then have "a div 2 ^ n * 2 ^ n < a" using sz anz apply (subst word_less_nat_alt) apply (subst unat_word_ariths) apply (subst unat_div) apply simp apply (rule order_le_less_trans [OF mod_less_eq_dividend]) apply (erule order_le_neq_trans [OF div_mult_le]) done also from xk le have "\ \ of_nat k * 2 ^ n" by (simp add: field_simps) finally show ?thesis using sz kv apply - apply (erule word_mult_less_dest [OF _ _ kn]) apply (simp add: unat_div) apply (rule order_le_less_trans [OF div_mult_le]) apply (rule unat_lt2p) done qed lemma nat_mod_power_lem: fixes a :: nat shows "1 < a \ a ^ n mod a ^ m = (if m \ n then 0 else a ^ n)" apply (clarsimp) apply (clarsimp simp add: le_iff_add power_add) done lemma power_mod_div: fixes x :: "nat" shows "x mod 2 ^ n div 2 ^ m = x div 2 ^ m mod 2 ^ (n - m)" (is "?LHS = ?RHS") proof (cases "n \ m") case True then have "?LHS = 0" apply - apply (rule div_less) apply (rule order_less_le_trans [OF mod_less_divisor]) apply simp apply simp done also have "\ = ?RHS" using True by simp finally show ?thesis . next case False then have lt: "m < n" by simp then obtain q where nv: "n = m + q" and "0 < q" by (auto dest: less_imp_Suc_add) then have "x mod 2 ^ n = 2 ^ m * (x div 2 ^ m mod 2 ^ q) + x mod 2 ^ m" by (simp add: power_add mod_mult2_eq) then have "?LHS = x div 2 ^ m mod 2 ^ q" by (simp add: div_add1_eq) also have "\ = ?RHS" using nv by simp finally show ?thesis . qed lemma word_power_mod_div: fixes x :: "'a::len word" shows "\ n < LENGTH('a); m < LENGTH('a)\ \ x mod 2 ^ n div 2 ^ m = x div 2 ^ m mod 2 ^ (n - m)" apply (simp add: word_arith_nat_div unat_mod power_mod_div) apply (subst unat_arith_simps(3)) apply (subst unat_mod) apply (subst unat_of_nat)+ apply (simp add: mod_mod_power min.commute) done lemma word_range_minus_1': fixes a :: "'a :: len word" shows "a \ 0 \ {a - 1<..b} = {a..b}" by (simp add: greaterThanAtMost_def atLeastAtMost_def greaterThan_def atLeast_def less_1_simp) lemma word_range_minus_1: fixes a :: "'a :: len word" shows "b \ 0 \ {a..b - 1} = {a.. 'b :: len word) x" by (simp add: ucast_def word_of_int_nat unat_def) lemmas if_fun_split = if_apply_def2 lemma i_hate_words_helper: "i \ (j - k :: nat) \ i \ j" by simp lemma i_hate_words: "unat (a :: 'a word) \ unat (b :: 'a :: len word) - Suc 0 \ a \ -1" apply (frule i_hate_words_helper) apply (subst(asm) word_le_nat_alt[symmetric]) apply (clarsimp simp only: word_minus_one_le) apply (simp only: linorder_not_less[symmetric]) apply (erule notE) apply (rule diff_Suc_less) apply (subst neq0_conv[symmetric]) apply (subst unat_eq_0) apply (rule notI, drule arg_cong[where f="(+) 1"]) apply simp done lemma overflow_plus_one_self: "(1 + p \ p) = (p = (-1 :: 'a :: len word))" apply (safe, simp_all) apply (rule ccontr) apply (drule plus_one_helper2) apply (rule notI) apply (drule arg_cong[where f="\x. x - 1"]) apply simp apply (simp add: field_simps) done lemma plus_1_less: "(x + 1 \ (x :: 'a :: len word)) = (x = -1)" apply (rule iffI) apply (rule ccontr) apply (cut_tac plus_one_helper2[where x=x, OF order_refl]) apply simp apply clarsimp apply (drule arg_cong[where f="\x. x - 1"]) apply simp apply simp done lemma pos_mult_pos_ge: "[|x > (0::int); n>=0 |] ==> n * x >= n*1" apply (simp only: mult_left_mono) done lemma If_eq_obvious: "x \ z \ ((if P then x else y) = z) = (\ P \ y = z)" by simp lemma Some_to_the: "v = Some x \ x = the v" by simp lemma dom_if_Some: "dom (\x. if P x then Some (f x) else g x) = {x. P x} \ dom g" by fastforce lemma dom_insert_absorb: "x \ dom f \ insert x (dom f) = dom f" by auto lemma emptyE2: "\ S = {}; x \ S \ \ P" by simp lemma mod_div_equality_div_eq: "a div b * b = (a - (a mod b) :: int)" by (simp add: field_simps) lemma zmod_helper: "n mod m = k \ ((n :: int) + a) mod m = (k + a) mod m" by (metis add.commute mod_add_right_eq) lemma int_div_sub_1: "\ m \ 1 \ \ (n - (1 :: int)) div m = (if m dvd n then (n div m) - 1 else n div m)" apply (subgoal_tac "m = 0 \ (n - (1 :: int)) div m = (if m dvd n then (n div m) - 1 else n div m)") apply fastforce apply (subst mult_cancel_right[symmetric]) apply (simp only: left_diff_distrib split: if_split) apply (simp only: mod_div_equality_div_eq) apply (clarsimp simp: field_simps) apply (clarsimp simp: dvd_eq_mod_eq_0) apply (cases "m = 1") apply simp apply (subst mod_diff_eq[symmetric], simp add: zmod_minus1) apply clarsimp apply (subst diff_add_cancel[where b=1, symmetric]) apply (subst mod_add_eq[symmetric]) apply (simp add: field_simps) apply (rule mod_pos_pos_trivial) apply (subst add_0_right[where a=0, symmetric]) apply (rule add_mono) apply simp apply simp apply (cases "(n - 1) mod m = m - 1") apply (drule zmod_helper[where a=1]) apply simp apply (subgoal_tac "1 + (n - 1) mod m \ m") apply simp apply (subst field_simps, rule zless_imp_add1_zle) apply simp done lemma ptr_add_image_multI: "\ \x y. (x * val = y * val') = (x * val'' = y); x * val'' \ S \ \ ptr_add ptr (x * val) \ (\p. ptr_add ptr (p * val')) ` S" apply (simp add: image_def) apply (erule rev_bexI) apply (rule arg_cong[where f="ptr_add ptr"]) apply simp done lemma shift_times_fold: "(x :: 'a :: len word) * (2 ^ n) << m = x << (m + n)" by (simp add: shiftl_t2n ac_simps power_add) lemma word_plus_strict_mono_right: fixes x :: "'a :: len word" shows "\y < z; x \ x + z\ \ x + y < x + z" by unat_arith lemma replicate_minus: "k < n \ replicate n False = replicate (n - k) False @ replicate k False" by (subst replicate_add [symmetric]) simp lemmas map_prod_split_imageI' = map_prod_imageI[where f="case_prod f" and g="case_prod g" and a="(a, b)" and b="(c, d)" for a b c d f g] lemmas map_prod_split_imageI = map_prod_split_imageI'[simplified] lemma word_div_mult: "0 < c \ a < b * c \ a div c < b" for a b c :: "'a::len word" by (rule classical) (use div_to_mult_word_lt [of b a c] in \auto simp add: word_less_nat_alt word_le_nat_alt unat_div\) lemma word_less_power_trans_ofnat: "\n < 2 ^ (m - k); k \ m; m < LENGTH('a)\ \ of_nat n * 2 ^ k < (2::'a::len word) ^ m" apply (subst mult.commute) apply (rule word_less_power_trans) apply (simp add: word_less_nat_alt) apply (subst unat_of_nat_eq) apply (erule order_less_trans) apply simp+ done lemma word_1_le_power: "n < LENGTH('a) \ (1 :: 'a :: len word) \ 2 ^ n" by (rule inc_le[where i=0, simplified], erule iffD2[OF p2_gt_0]) lemma enum_word_div: fixes v :: "'a :: len word" shows "\xs ys. enum = xs @ [v] @ ys \ (\x \ set xs. x < v) \ (\y \ set ys. v < y)" apply (simp only: enum_word_def) apply (subst upt_add_eq_append'[where j="unat v"]) apply simp apply (rule order_less_imp_le, simp) apply (simp add: upt_conv_Cons) apply (intro exI conjI) apply fastforce apply clarsimp apply (drule of_nat_mono_maybe[rotated, where 'a='a]) apply simp apply simp apply (clarsimp simp: Suc_le_eq) apply (drule of_nat_mono_maybe[rotated, where 'a='a]) apply simp apply simp done lemma of_bool_nth: "of_bool (x !! v) = (x >> v) && 1" apply (rule word_eqI) apply (simp add: nth_shiftr cong: rev_conj_cong) done lemma unat_1_0: "1 \ (x::'a::len word) = (0 < unat x)" by (auto simp add: word_le_nat_alt) lemma x_less_2_0_1': fixes x :: "'a::len word" shows "\LENGTH('a) \ 1; x < 2\ \ x = 0 \ x = 1" apply (induct x) apply clarsimp+ by (metis Suc_eq_plus1 add_lessD1 less_irrefl one_add_one unatSuc word_less_nat_alt) lemmas word_add_le_iff2 = word_add_le_iff [folded no_olen_add_nat] lemma of_nat_power: shows "\ p < 2 ^ x; x < len_of TYPE ('a) \ \ of_nat p < (2 :: 'a :: len word) ^ x" apply (rule order_less_le_trans) apply (rule of_nat_mono_maybe) apply (erule power_strict_increasing) apply simp apply assumption apply (simp add: word_unat_power) done lemma of_nat_n_less_equal_power_2: "n < LENGTH('a::len) \ ((of_nat n)::'a word) < 2 ^ n" apply (induct n) apply clarsimp apply clarsimp apply (metis of_nat_power n_less_equal_power_2 of_nat_Suc power_Suc) done lemma eq_mask_less: fixes w :: "'a::len word" assumes eqm: "w = w && mask n" and sz: "n < len_of TYPE ('a)" shows "w < (2::'a word) ^ n" by (subst eqm, rule and_mask_less' [OF sz]) lemma of_nat_mono_maybe': fixes Y :: "nat" assumes xlt: "x < 2 ^ len_of TYPE ('a)" assumes ylt: "y < 2 ^ len_of TYPE ('a)" shows "(y < x) = (of_nat y < (of_nat x :: 'a :: len word))" apply (subst word_less_nat_alt) apply (subst unat_of_nat)+ apply (subst mod_less) apply (rule ylt) apply (subst mod_less) apply (rule xlt) apply simp done lemma shiftr_mask_eq: fixes x :: "'a :: len word" shows "(x >> n) && mask (size x - n) = x >> n" apply (rule word_eqI) apply (clarsimp simp: word_size nth_shiftr) apply (rule iffI) apply clarsimp apply (clarsimp) apply (drule test_bit_size) apply (simp add: word_size) done lemma shiftr_mask_eq': fixes x :: "'a :: len word" shows "m = (size x - n) \ (x >> n) && mask m = x >> n" by (simp add: shiftr_mask_eq) lemma bang_big: "n \ size (x::'a::len0 word) \ (x !! n) = False" by (simp add: test_bit_bl word_size) lemma bang_conj_lt: fixes x :: "'a :: len word" shows "(x !! m \ m < LENGTH('a)) = x !! m" apply (cases "m < LENGTH('a)") apply simp apply (simp add: not_less bang_big word_size) done lemma dom_if: "dom (\a. if a \ addrs then Some (f a) else g a) = addrs \ dom g" by (auto simp: dom_def split: if_split) lemma less_is_non_zero_p1: fixes a :: "'a :: len word" shows "a < k \ a + 1 \ 0" apply (erule contrapos_pn) apply (drule max_word_wrap) apply (simp add: not_less) done lemma of_nat_mono_maybe_le: "\x < 2 ^ LENGTH('a); y < 2 ^ LENGTH('a)\ \ (y \ x) = ((of_nat y :: 'a :: len word) \ of_nat x)" apply (clarsimp simp: le_less) apply (rule disj_cong) apply (rule of_nat_mono_maybe', assumption+) apply (simp add: word_unat.norm_eq_iff [symmetric]) done lemma neg_mask_bang: "(~~ mask n :: 'a :: len word) !! m = (n \ m \ m < LENGTH('a))" apply (cases "m < LENGTH('a)") apply (simp add: word_ops_nth_size word_size not_less) apply (simp add: not_less bang_big word_size) done lemma mask_AND_NOT_mask: "(w && ~~ mask n) && mask n = 0" by (rule word_eqI) (clarsimp simp add: word_size neg_mask_bang) lemma AND_NOT_mask_plus_AND_mask_eq: "(w && ~~ mask n) + (w && mask n) = w" apply (rule word_eqI) apply (rename_tac m) apply (simp add: word_size) apply (cut_tac word_plus_and_or_coroll[of "w && ~~ mask n" "w && mask n"]) apply (simp add: word_ao_dist2[symmetric] word_size neg_mask_bang) apply (rule word_eqI) apply (rename_tac m) apply (simp add: word_size neg_mask_bang) done lemma mask_eqI: fixes x :: "'a :: len word" assumes m1: "x && mask n = y && mask n" and m2: "x && ~~ mask n = y && ~~ mask n" shows "x = y" proof (subst bang_eq, rule allI) fix m show "x !! m = y !! m" proof (cases "m < n") case True then have "x !! m = ((x && mask n) !! m)" by (simp add: word_size bang_conj_lt) also have "\ = ((y && mask n) !! m)" using m1 by simp also have "\ = y !! m" using True by (simp add: word_size bang_conj_lt) finally show ?thesis . next case False then have "x !! m = ((x && ~~ mask n) !! m)" by (simp add: neg_mask_bang bang_conj_lt) also have "\ = ((y && ~~ mask n) !! m)" using m2 by simp also have "\ = y !! m" using False by (simp add: neg_mask_bang bang_conj_lt) finally show ?thesis . qed qed lemma nat_less_power_trans2: fixes n :: nat shows "\n < 2 ^ (m - k); k \ m\ \ n * 2 ^ k < 2 ^ m" by (subst mult.commute, erule (1) nat_less_power_trans) lemma nat_move_sub_le: "(a::nat) + b \ c \ a \ c - b" by arith lemma neq_0_no_wrap: fixes x :: "'a :: len word" shows "\ x \ x + y; x \ 0 \ \ x + y \ 0" by clarsimp lemma plus_minus_one_rewrite: "v + (- 1 :: ('a :: {ring, one, uminus})) \ v - 1" by (simp add: field_simps) lemma power_minus_is_div: "b \ a \ (2 :: nat) ^ (a - b) = 2 ^ a div 2 ^ b" apply (induct a arbitrary: b) apply simp apply (erule le_SucE) apply (clarsimp simp:Suc_diff_le le_iff_add power_add) apply simp done lemma two_pow_div_gt_le: "v < 2 ^ n div (2 ^ m :: nat) \ m \ n" by (clarsimp dest!: less_two_pow_divD) lemma unatSuc2: fixes n :: "'a :: len word" shows "n + 1 \ 0 \ unat (n + 1) = Suc (unat n)" by (simp add: add.commute unatSuc) lemma word_of_nat_less: "\ n < unat x \ \ of_nat n < x" apply (simp add: word_less_nat_alt) apply (erule order_le_less_trans[rotated]) apply (simp add: unat_of_nat) done lemma word_of_nat_le: "n \ unat x \ of_nat n \ x" apply (simp add: word_le_nat_alt unat_of_nat) apply (erule order_trans[rotated]) apply simp done lemma word_unat_less_le: "a \ of_nat b \ unat a \ b" by (metis eq_iff le_cases le_unat_uoi word_of_nat_le) lemma and_eq_0_is_nth: fixes x :: "'a :: len word" shows "y = 1 << n \ ((x && y) = 0) = (\ (x !! n))" apply safe apply (drule_tac u="(x && (1 << n))" and x=n in word_eqD) apply (simp add: nth_w2p) apply (simp add: test_bit_bin) apply (rule word_eqI) apply (simp add: nth_w2p) done lemmas arg_cong_Not = arg_cong [where f=Not] lemmas and_neq_0_is_nth = arg_cong_Not [OF and_eq_0_is_nth, simplified] lemma nth_is_and_neq_0: "(x::'a::len word) !! n = (x && 2 ^ n \ 0)" by (subst and_neq_0_is_nth; rule refl) lemma mask_Suc_0 : "mask (Suc 0) = 1" by (simp add: mask_def) lemma ucast_ucast_add: fixes x :: "'a :: len word" fixes y :: "'b :: len word" shows "LENGTH('b) \ LENGTH('a) \ ucast (ucast x + y) = x + ucast y" apply (rule word_unat.Rep_eqD) apply (simp add: unat_ucast unat_word_ariths mod_mod_power min.absorb2 unat_of_nat) apply (subst mod_add_left_eq[symmetric]) apply (simp add: mod_mod_power min.absorb2) apply (subst mod_add_right_eq) apply simp done lemma word_shift_zero: "\ x << n = 0; x \ 2^m; m + n < LENGTH('a)\ \ (x::'a::len word) = 0" apply (rule ccontr) apply (drule (2) word_shift_nonzero) apply simp done lemma neg_mask_mono_le: "(x :: 'a :: len word) \ y \ x && ~~ mask n \ y && ~~ mask n" proof (rule ccontr, simp add: linorder_not_le, cases "n < LENGTH('a)") case False show "y && ~~ mask n < x && ~~ mask n \ False" using False by (simp add: mask_def linorder_not_less power_overflow) next case True assume a: "x \ y" and b: "y && ~~ mask n < x && ~~ mask n" have word_bits: "n < LENGTH('a)" using True by assumption have "y \ (y && ~~ mask n) + (y && mask n)" by (simp add: word_plus_and_or_coroll2 add.commute) also have "\ \ (y && ~~ mask n) + 2 ^ n" apply (rule word_plus_mono_right) apply (rule order_less_imp_le, rule and_mask_less_size) apply (simp add: word_size word_bits) apply (rule is_aligned_no_overflow'', simp_all add: is_aligned_neg_mask word_bits) apply (rule not_greatest_aligned, rule b) apply (simp_all add: is_aligned_neg_mask) done also have "\ \ x && ~~ mask n" using b apply - apply (subst add.commute, rule le_plus) apply (rule aligned_at_least_t2n_diff, simp_all add: is_aligned_neg_mask) apply (rule ccontr, simp add: linorder_not_le) apply (drule aligned_small_is_0[rotated], simp_all add: is_aligned_neg_mask) done also have "\ \ x" by (rule word_and_le2) also have "x \ y" by fact finally show "False" using b by simp qed lemma bool_mask': fixes x :: "'a :: len word" shows "2 < LENGTH('a) \ (0 < x && 1) = (x && 1 = 1)" apply (rule iffI) prefer 2 apply simp apply (subgoal_tac "x && mask 1 < 2^1") prefer 2 apply (rule and_mask_less_size) apply (simp add: word_size) apply (simp add: mask_def) apply (drule word_less_cases [where y=2]) apply (erule disjE, simp) apply simp done lemma sint_eq_uint: "\ msb x \ sint x = uint x" apply (rule word_uint.Abs_eqD, subst word_sint.Rep_inverse) apply simp_all apply (cut_tac x=x in word_sint.Rep) apply (clarsimp simp add: uints_num sints_num) apply (rule conjI) apply (rule ccontr) apply (simp add: linorder_not_le word_msb_sint[symmetric]) apply (erule order_less_le_trans) apply simp done lemma scast_eq_ucast: "\ msb x \ scast x = ucast x" by (simp add: scast_def ucast_def sint_eq_uint) lemma lt1_neq0: fixes x :: "'a :: len word" shows "(1 \ x) = (x \ 0)" by unat_arith lemma word_plus_one_nonzero: fixes x :: "'a :: len word" shows "\x \ x + y; y \ 0\ \ x + 1 \ 0" apply (subst lt1_neq0 [symmetric]) apply (subst olen_add_eqv [symmetric]) apply (erule word_random) apply (simp add: lt1_neq0) done lemma word_sub_plus_one_nonzero: fixes n :: "'a :: len word" shows "\n' \ n; n' \ 0\ \ (n - n') + 1 \ 0" apply (subst lt1_neq0 [symmetric]) apply (subst olen_add_eqv [symmetric]) apply (rule word_random [where x' = n']) apply simp apply (erule word_sub_le) apply (simp add: lt1_neq0) done lemma word_le_minus_mono_right: fixes x :: "'a :: len word" shows "\ z \ y; y \ x; z \ x \ \ x - y \ x - z" apply (rule word_sub_mono) apply simp apply assumption apply (erule word_sub_le) apply (erule word_sub_le) done lemma drop_append_miracle: "n = length xs \ drop n (xs @ ys) = ys" by simp lemma foldr_does_nothing_to_xf: "\ \x s. x \ set xs \ xf (f x s) = xf s \ \ xf (foldr f xs s) = xf s" by (induct xs, simp_all) lemma nat_less_mult_monoish: "\ a < b; c < (d :: nat) \ \ (a + 1) * (c + 1) <= b * d" apply (drule Suc_leI)+ apply (drule(1) mult_le_mono) apply simp done lemma word_0_sle_from_less[unfolded word_size]: "\ x < 2 ^ (size x - 1) \ \ 0 <=s x" apply (clarsimp simp: word_sle_msb_le) apply (simp add: word_msb_nth) apply (subst (asm) word_test_bit_def [symmetric]) apply (drule less_mask_eq) apply (drule_tac x="size x - 1" in word_eqD) apply (simp add: word_size) done lemma not_msb_from_less: "(v :: 'a word) < 2 ^ (LENGTH('a :: len) - 1) \ \ msb v" apply (clarsimp simp add: msb_nth) apply (drule less_mask_eq) apply (drule word_eqD, drule(1) iffD2) apply simp done lemma distinct_lemma: "f x \ f y \ x \ y" by auto lemma ucast_sub_ucast: fixes x :: "'a::len word" assumes "y \ x" assumes T: "LENGTH('a) \ LENGTH('b)" shows "ucast (x - y) = (ucast x - ucast y :: 'b::len word)" proof - from T have P: "unat x < 2 ^ LENGTH('b)" "unat y < 2 ^ LENGTH('b)" by (fastforce intro!: less_le_trans[OF unat_lt2p])+ then show ?thesis by (simp add: unat_arith_simps unat_ucast assms[simplified unat_arith_simps]) qed lemma word_1_0: "\a + (1::('a::len) word) \ b; a < of_nat x\ \ a < b" by unat_arith lemma unat_of_nat_less:"\ a < b; unat b = c \ \ a < of_nat c" by fastforce lemma word_le_plus_1: "\ (y::('a::len) word) < y + n; a < n \ \ y + a \ y + a + 1" by unat_arith lemma word_le_plus:"\(a::('a::len) word) < a + b; c < b\ \ a \ a + c" by (metis order_less_imp_le word_random) (* * Basic signed arithemetic properties. *) lemma sint_minus1 [simp]: "(sint x = -1) = (x = -1)" by (metis sint_n1 word_sint.Rep_inverse') lemma sint_0 [simp]: "(sint x = 0) = (x = 0)" by (metis sint_0 word_sint.Rep_inverse') (* It is not always that case that "sint 1 = 1", because of 1-bit word sizes. * This lemma produces the different cases. *) lemma sint_1_cases: "\ \ len_of TYPE ('a::len) = 1; (a::'a word) = 0; sint a = 0 \ \ P; \ len_of TYPE ('a) = 1; a = 1; sint (1 :: 'a word) = -1 \ \ P; \ len_of TYPE ('a) > 1; sint (1 :: 'a word) = 1 \ \ P \ \ P" apply atomize_elim apply (case_tac "len_of TYPE ('a) = 1") apply clarsimp apply (subgoal_tac "(UNIV :: 'a word set) = {0, 1}") apply (metis UNIV_I insert_iff singletonE) apply (subst word_unat.univ) apply (clarsimp simp: unats_def image_def) apply (rule set_eqI, rule iffI) apply clarsimp apply (metis One_nat_def less_2_cases of_nat_1 semiring_1_class.of_nat_0) apply clarsimp apply (metis Abs_fnat_hom_0 Suc_1 lessI of_nat_1 zero_less_Suc) apply clarsimp apply (metis One_nat_def arith_is_1 le_def len_gt_0) done lemma sint_int_min: "sint (- (2 ^ (LENGTH('a) - Suc 0)) :: ('a::len) word) = - (2 ^ (LENGTH('a) - Suc 0))" apply (subst word_sint.Abs_inverse' [where r="- (2 ^ (LENGTH('a) - Suc 0))"]) apply (clarsimp simp: sints_num) apply (clarsimp simp: wi_hom_syms word_of_int_2p) apply clarsimp done lemma sint_int_max_plus_1: "sint (2 ^ (LENGTH('a) - Suc 0) :: ('a::len) word) = - (2 ^ (LENGTH('a) - Suc 0))" apply (subst word_of_int_2p [symmetric]) apply (subst int_word_sint) apply clarsimp apply (metis Suc_pred int_word_uint len_gt_0 power_Suc uint_eq_0 word_of_int_2p word_pow_0) done lemma sbintrunc_eq_in_range: "(sbintrunc n x = x) = (x \ range (sbintrunc n))" "(x = sbintrunc n x) = (x \ range (sbintrunc n))" apply (simp_all add: image_def) apply (metis sbintrunc_sbintrunc)+ done lemma sbintrunc_If: "- 3 * (2 ^ n) \ x \ x < 3 * (2 ^ n) \ sbintrunc n x = (if x < - (2 ^ n) then x + 2 * (2 ^ n) else if x \ 2 ^ n then x - 2 * (2 ^ n) else x)" apply (simp add: no_sbintr_alt2, safe) apply (simp add: mod_pos_geq) apply (subst mod_add_self1[symmetric], simp) done lemma signed_arith_eq_checks_to_ord: "(sint a + sint b = sint (a + b )) = ((a <=s a + b) = (0 <=s b))" "(sint a - sint b = sint (a - b )) = ((0 <=s a - b) = (b <=s a))" "(- sint a = sint (- a)) = (0 <=s (- a) = (a <=s 0))" using sint_range'[where x=a] sint_range'[where x=b] by (simp_all add: sint_word_ariths word_sle_def word_sless_alt sbintrunc_If) (* Basic proofs that signed word div/mod operations are * truncations of their integer counterparts. *) lemma signed_div_arith: "sint ((a::('a::len) word) sdiv b) = sbintrunc (LENGTH('a) - 1) (sint a sdiv sint b)" apply (subst word_sbin.norm_Rep [symmetric]) apply (subst bin_sbin_eq_iff' [symmetric]) apply simp apply (subst uint_sint [symmetric]) apply (clarsimp simp: sdiv_int_def sdiv_word_def) apply (metis word_ubin.eq_norm) done lemma signed_mod_arith: "sint ((a::('a::len) word) smod b) = sbintrunc (LENGTH('a) - 1) (sint a smod sint b)" apply (subst word_sbin.norm_Rep [symmetric]) apply (subst bin_sbin_eq_iff' [symmetric]) apply simp apply (subst uint_sint [symmetric]) apply (clarsimp simp: smod_int_def smod_word_def) apply (metis word_ubin.eq_norm) done (* Signed word arithmetic overflow constraints. *) lemma signed_arith_ineq_checks_to_eq: "((- (2 ^ (size a - 1)) \ (sint a + sint b)) \ (sint a + sint b \ (2 ^ (size a - 1) - 1))) = (sint a + sint b = sint (a + b ))" "((- (2 ^ (size a - 1)) \ (sint a - sint b)) \ (sint a - sint b \ (2 ^ (size a - 1) - 1))) = (sint a - sint b = sint (a - b))" "((- (2 ^ (size a - 1)) \ (- sint a)) \ (- sint a) \ (2 ^ (size a - 1) - 1)) = ((- sint a) = sint (- a))" "((- (2 ^ (size a - 1)) \ (sint a * sint b)) \ (sint a * sint b \ (2 ^ (size a - 1) - 1))) = (sint a * sint b = sint (a * b))" "((- (2 ^ (size a - 1)) \ (sint a sdiv sint b)) \ (sint a sdiv sint b \ (2 ^ (size a - 1) - 1))) = (sint a sdiv sint b = sint (a sdiv b))" "((- (2 ^ (size a - 1)) \ (sint a smod sint b)) \ (sint a smod sint b \ (2 ^ (size a - 1) - 1))) = (sint a smod sint b = sint (a smod b))" by (auto simp: sint_word_ariths word_size signed_div_arith signed_mod_arith sbintrunc_eq_in_range range_sbintrunc) lemma signed_arith_sint: "((- (2 ^ (size a - 1)) \ (sint a + sint b)) \ (sint a + sint b \ (2 ^ (size a - 1) - 1))) \ sint (a + b) = (sint a + sint b)" "((- (2 ^ (size a - 1)) \ (sint a - sint b)) \ (sint a - sint b \ (2 ^ (size a - 1) - 1))) \ sint (a - b) = (sint a - sint b)" "((- (2 ^ (size a - 1)) \ (- sint a)) \ (- sint a) \ (2 ^ (size a - 1) - 1)) \ sint (- a) = (- sint a)" "((- (2 ^ (size a - 1)) \ (sint a * sint b)) \ (sint a * sint b \ (2 ^ (size a - 1) - 1))) \ sint (a * b) = (sint a * sint b)" "((- (2 ^ (size a - 1)) \ (sint a sdiv sint b)) \ (sint a sdiv sint b \ (2 ^ (size a - 1) - 1))) \ sint (a sdiv b) = (sint a sdiv sint b)" "((- (2 ^ (size a - 1)) \ (sint a smod sint b)) \ (sint a smod sint b \ (2 ^ (size a - 1) - 1))) \ sint (a smod b) = (sint a smod sint b)" by (subst (asm) signed_arith_ineq_checks_to_eq; simp)+ lemma signed_mult_eq_checks_double_size: assumes mult_le: "(2 ^ (len_of TYPE ('a) - 1) + 1) ^ 2 \ (2 :: int) ^ (len_of TYPE ('b) - 1)" and le: "2 ^ (LENGTH('a) - 1) \ (2 :: int) ^ (len_of TYPE ('b) - 1)" shows "(sint (a :: 'a :: len word) * sint b = sint (a * b)) = (scast a * scast b = (scast (a * b) :: 'b :: len word))" proof - have P: "sbintrunc (size a - 1) (sint a * sint b) \ range (sbintrunc (size a - 1))" by simp have abs: "!! x :: 'a word. abs (sint x) < 2 ^ (size a - 1) + 1" apply (cut_tac x=x in sint_range') apply (simp add: abs_le_iff word_size) done have abs_ab: "abs (sint a * sint b) < 2 ^ (LENGTH('b) - 1)" using abs_mult_less[OF abs[where x=a] abs[where x=b]] mult_le by (simp add: abs_mult power2_eq_square word_size) show ?thesis using P[unfolded range_sbintrunc] abs_ab le apply (simp add: sint_word_ariths scast_def) apply (simp add: wi_hom_mult) apply (subst word_sint.Abs_inject, simp_all) apply (simp add: sints_def range_sbintrunc abs_less_iff) apply clarsimp apply (simp add: sints_def range_sbintrunc word_size) apply (auto elim: order_less_le_trans order_trans[rotated]) done qed (* Properties about signed division. *) lemma int_sdiv_simps [simp]: "(a :: int) sdiv 1 = a" "(a :: int) sdiv 0 = 0" "(a :: int) sdiv -1 = -a" apply (auto simp: sdiv_int_def sgn_if) done lemma sgn_div_eq_sgn_mult: "a div b \ 0 \ sgn ((a :: int) div b) = sgn (a * b)" apply (clarsimp simp: sgn_if zero_le_mult_iff neg_imp_zdiv_nonneg_iff not_less) apply (metis less_le mult_le_0_iff neg_imp_zdiv_neg_iff not_less pos_imp_zdiv_neg_iff zdiv_eq_0_iff) done lemma sgn_sdiv_eq_sgn_mult: "a sdiv b \ 0 \ sgn ((a :: int) sdiv b) = sgn (a * b)" by (auto simp: sdiv_int_def sgn_div_eq_sgn_mult sgn_mult) lemma int_sdiv_same_is_1 [simp]: "a \ 0 \ ((a :: int) sdiv b = a) = (b = 1)" apply (rule iffI) apply (clarsimp simp: sdiv_int_def) apply (subgoal_tac "b > 0") apply (case_tac "a > 0") apply (clarsimp simp: sgn_if) - apply (clarsimp simp: sign_simps not_less) + apply (clarsimp simp: algebra_split_simps not_less) apply (metis int_div_same_is_1 le_neq_trans minus_minus neg_0_le_iff_le neg_equal_0_iff_equal) apply (case_tac "a > 0") apply (case_tac "b = 0") apply clarsimp apply (rule classical) apply (clarsimp simp: sgn_mult not_less) apply (metis le_less neg_0_less_iff_less not_less_iff_gr_or_eq pos_imp_zdiv_neg_iff) apply (rule classical) - apply (clarsimp simp: sign_simps sgn_mult not_less sgn_if split: if_splits) + apply (clarsimp simp: algebra_split_simps sgn_mult not_less sgn_if split: if_splits) apply (metis antisym less_le neg_imp_zdiv_nonneg_iff) apply (clarsimp simp: sdiv_int_def sgn_if) done lemma int_sdiv_negated_is_minus1 [simp]: "a \ 0 \ ((a :: int) sdiv b = - a) = (b = -1)" apply (clarsimp simp: sdiv_int_def) apply (rule iffI) apply (subgoal_tac "b < 0") apply (case_tac "a > 0") - apply (clarsimp simp: sgn_if sign_simps not_less) + apply (clarsimp simp: sgn_if algebra_split_simps not_less) apply (case_tac "sgn (a * b) = -1") - apply (clarsimp simp: not_less sign_simps) - apply (clarsimp simp: sign_simps not_less) + apply (clarsimp simp: not_less algebra_split_simps) + apply (clarsimp simp: algebra_split_simps not_less) apply (rule classical) apply (case_tac "b = 0") apply (clarsimp simp: not_less sgn_mult) apply (case_tac "a > 0") apply (clarsimp simp: not_less sgn_mult) apply (metis less_le neg_less_0_iff_less not_less_iff_gr_or_eq pos_imp_zdiv_neg_iff) apply (clarsimp simp: not_less sgn_mult) apply (metis antisym_conv div_minus_right neg_imp_zdiv_nonneg_iff neg_le_0_iff_le not_less) apply (clarsimp simp: sgn_if) done lemma sdiv_int_range: "(a :: int) sdiv b \ { - (abs a) .. (abs a) }" apply (unfold sdiv_int_def) apply (subgoal_tac "(abs a) div (abs b) \ (abs a)") apply (clarsimp simp: sgn_if) apply (meson abs_ge_zero neg_le_0_iff_le nonneg_mod_div order_trans) apply (metis abs_eq_0 abs_ge_zero div_by_0 zdiv_le_dividend zero_less_abs_iff) done lemma word_sdiv_div1 [simp]: "(a :: ('a::len) word) sdiv 1 = a" apply (rule sint_1_cases [where a=a]) apply (clarsimp simp: sdiv_word_def sdiv_int_def) apply (clarsimp simp: sdiv_word_def sdiv_int_def simp del: sint_minus1) apply (clarsimp simp: sdiv_word_def) done lemma sdiv_int_div_0 [simp]: "(x :: int) sdiv 0 = 0" by (clarsimp simp: sdiv_int_def) lemma sdiv_int_0_div [simp]: "0 sdiv (x :: int) = 0" by (clarsimp simp: sdiv_int_def) lemma word_sdiv_div0 [simp]: "(a :: ('a::len) word) sdiv 0 = 0" apply (auto simp: sdiv_word_def sdiv_int_def sgn_if) done lemma word_sdiv_div_minus1 [simp]: "(a :: ('a::len) word) sdiv -1 = -a" apply (auto simp: sdiv_word_def sdiv_int_def sgn_if) apply (metis wi_hom_neg word_sint.Rep_inverse') done lemmas word_sdiv_0 = word_sdiv_div0 lemma sdiv_word_min: "- (2 ^ (size a - 1)) \ sint (a :: ('a::len) word) sdiv sint (b :: ('a::len) word)" apply (clarsimp simp: word_size) apply (cut_tac sint_range' [where x=a]) apply (cut_tac sint_range' [where x=b]) apply clarsimp apply (insert sdiv_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: max_def abs_if split: if_split_asm) done lemma sdiv_word_max: "(sint (a :: ('a::len) word) sdiv sint (b :: ('a::len) word) < (2 ^ (size a - 1))) = ((a \ - (2 ^ (size a - 1)) \ (b \ -1)))" (is "?lhs = (\ ?a_int_min \ \ ?b_minus1)") proof (rule classical) assume not_thesis: "\ ?thesis" have not_zero: "b \ 0" using not_thesis by (clarsimp) have result_range: "sint a sdiv sint b \ (sints (size a)) \ {2 ^ (size a - 1)}" apply (cut_tac sdiv_int_range [where a="sint a" and b="sint b"]) apply (erule rev_subsetD) using sint_range' [where x=a] sint_range' [where x=b] apply (auto simp: max_def abs_if word_size sints_num) done have result_range_overflow: "(sint a sdiv sint b = 2 ^ (size a - 1)) = (?a_int_min \ ?b_minus1)" apply (rule iffI [rotated]) apply (clarsimp simp: sdiv_int_def sgn_if word_size sint_int_min) apply (rule classical) apply (case_tac "?a_int_min") apply (clarsimp simp: word_size sint_int_min) apply (metis diff_0_right int_sdiv_negated_is_minus1 minus_diff_eq minus_int_code(2) power_eq_0_iff sint_minus1 zero_neq_numeral) apply (subgoal_tac "abs (sint a) < 2 ^ (size a - 1)") apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1] apply (clarsimp simp: word_size) apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1] apply (insert word_sint.Rep [where x="a"])[1] apply (clarsimp simp: minus_le_iff word_size abs_if sints_num split: if_split_asm) apply (metis minus_minus sint_int_min word_sint.Rep_inject) done have result_range_simple: "(sint a sdiv sint b \ (sints (size a))) \ ?thesis" apply (insert sdiv_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: word_size sints_num sint_int_min) done show ?thesis apply (rule UnE [OF result_range result_range_simple]) apply simp apply (clarsimp simp: word_size) using result_range_overflow apply (clarsimp simp: word_size) done qed lemmas sdiv_word_min' = sdiv_word_min [simplified word_size, simplified] lemmas sdiv_word_max' = sdiv_word_max [simplified word_size, simplified] lemmas word_sdiv_numerals_lhs = sdiv_word_def[where a="numeral x" for x] sdiv_word_def[where a=0] sdiv_word_def[where a=1] lemmas word_sdiv_numerals = word_sdiv_numerals_lhs[where b="numeral y" for y] word_sdiv_numerals_lhs[where b=0] word_sdiv_numerals_lhs[where b=1] (* * Signed modulo properties. *) lemma smod_int_alt_def: "(a::int) smod b = sgn (a) * (abs a mod abs b)" apply (clarsimp simp: smod_int_def sdiv_int_def) - apply (clarsimp simp: minus_div_mult_eq_mod [symmetric] abs_sgn sgn_mult sgn_if sign_simps) + apply (clarsimp simp: minus_div_mult_eq_mod [symmetric] abs_sgn sgn_mult sgn_if algebra_split_simps) done lemma smod_int_range: "b \ 0 \ (a::int) smod b \ { - abs b + 1 .. abs b - 1 }" apply (case_tac "b > 0") apply (insert pos_mod_conj [where a=a and b=b])[1] apply (insert pos_mod_conj [where a="-a" and b=b])[1] apply (auto simp: smod_int_alt_def algebra_simps sgn_if abs_if not_less add1_zle_eq [simplified add.commute])[1] apply (metis add_nonneg_nonneg int_one_le_iff_zero_less le_less less_add_same_cancel2 not_le pos_mod_conj) apply (metis (full_types) add.inverse_inverse eucl_rel_int eucl_rel_int_iff le_less_trans neg_0_le_iff_le) apply (insert neg_mod_conj [where a=a and b="b"])[1] apply (insert neg_mod_conj [where a="-a" and b="b"])[1] apply (clarsimp simp: smod_int_alt_def algebra_simps sgn_if abs_if not_less add1_zle_eq [simplified add.commute]) apply (metis neg_0_less_iff_less neg_mod_conj not_le not_less_iff_gr_or_eq order_trans pos_mod_conj) done lemma smod_int_compares: "\ 0 \ a; 0 < b \ \ (a :: int) smod b < b" "\ 0 \ a; 0 < b \ \ 0 \ (a :: int) smod b" "\ a \ 0; 0 < b \ \ -b < (a :: int) smod b" "\ a \ 0; 0 < b \ \ (a :: int) smod b \ 0" "\ 0 \ a; b < 0 \ \ (a :: int) smod b < - b" "\ 0 \ a; b < 0 \ \ 0 \ (a :: int) smod b" "\ a \ 0; b < 0 \ \ (a :: int) smod b \ 0" "\ a \ 0; b < 0 \ \ b \ (a :: int) smod b" apply (insert smod_int_range [where a=a and b=b]) apply (auto simp: add1_zle_eq smod_int_alt_def sgn_if) done lemma smod_int_mod_0 [simp]: "x smod (0 :: int) = x" by (clarsimp simp: smod_int_def) lemma smod_int_0_mod [simp]: "0 smod (x :: int) = 0" by (clarsimp simp: smod_int_alt_def) lemma smod_word_mod_0 [simp]: "x smod (0 :: ('a::len) word) = x" by (clarsimp simp: smod_word_def) lemma smod_word_0_mod [simp]: "0 smod (x :: ('a::len) word) = 0" by (clarsimp simp: smod_word_def) lemma smod_word_max: "sint (a::'a word) smod sint (b::'a word) < 2 ^ (LENGTH('a::len) - Suc 0)" apply (case_tac "b = 0") apply (insert word_sint.Rep [where x=a, simplified sints_num])[1] apply (clarsimp) apply (insert word_sint.Rep [where x="b", simplified sints_num])[1] apply (insert smod_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: abs_if split: if_split_asm) done lemma smod_word_min: "- (2 ^ (LENGTH('a::len) - Suc 0)) \ sint (a::'a word) smod sint (b::'a word)" apply (case_tac "b = 0") apply (insert word_sint.Rep [where x=a, simplified sints_num])[1] apply clarsimp apply (insert word_sint.Rep [where x=b, simplified sints_num])[1] apply (insert smod_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: abs_if add1_zle_eq split: if_split_asm) done lemma smod_word_alt_def: "(a :: ('a::len) word) smod b = a - (a sdiv b) * b" apply (case_tac "a \ - (2 ^ (LENGTH('a) - 1)) \ b \ -1") apply (clarsimp simp: smod_word_def sdiv_word_def smod_int_def minus_word.abs_eq [symmetric] times_word.abs_eq [symmetric]) apply (clarsimp simp: smod_word_def smod_int_def) done lemmas word_smod_numerals_lhs = smod_word_def[where a="numeral x" for x] smod_word_def[where a=0] smod_word_def[where a=1] lemmas word_smod_numerals = word_smod_numerals_lhs[where b="numeral y" for y] word_smod_numerals_lhs[where b=0] word_smod_numerals_lhs[where b=1] lemma sint_of_int_eq: "\ - (2 ^ (LENGTH('a) - 1)) \ x; x < 2 ^ (LENGTH('a) - 1) \ \ sint (of_int x :: ('a::len) word) = x" apply (clarsimp simp: word_of_int int_word_sint) apply (subst int_mod_eq') apply simp apply (subst (2) power_minus_simp) apply clarsimp apply clarsimp apply clarsimp done lemma of_int_sint [simp]: "of_int (sint a) = a" apply (insert word_sint.Rep [where x=a]) apply (clarsimp simp: word_of_int) done lemma nth_w2p_scast [simp]: "((scast ((2::'a::len signed word) ^ n) :: 'a word) !! m) \ ((((2::'a::len word) ^ n) :: 'a word) !! m)" apply (subst nth_w2p) apply (case_tac "n \ LENGTH('a)") apply (subst power_overflow, simp) apply clarsimp apply (metis nth_w2p scast_def bang_conj_lt len_signed nth_word_of_int word_sint.Rep_inverse) done lemma scast_2_power [simp]: "scast ((2 :: 'a::len signed word) ^ x) = ((2 :: 'a word) ^ x)" by (clarsimp simp: word_eq_iff) lemma scast_bit_test [simp]: "scast ((1 :: 'a::len signed word) << n) = (1 :: 'a word) << n" by (clarsimp simp: word_eq_iff) lemma ucast_nat_def': "of_nat (unat x) = (ucast :: 'a :: len word \ ('b :: len) signed word) x" by (simp add: ucast_def word_of_int_nat unat_def) lemma mod_mod_power_int: fixes k :: int shows "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ (min m n)" by (metis bintrunc_bintrunc_min bintrunc_mod2p min.commute) (* Normalise combinations of scast and ucast. *) lemma ucast_distrib: fixes M :: "'a::len word \ 'a::len word \ 'a::len word" fixes M' :: "'b::len word \ 'b::len word \ 'b::len word" fixes L :: "int \ int \ int" assumes lift_M: "\x y. uint (M x y) = L (uint x) (uint y) mod 2 ^ LENGTH('a)" assumes lift_M': "\x y. uint (M' x y) = L (uint x) (uint y) mod 2 ^ LENGTH('b)" assumes distrib: "\x y. (L (x mod (2 ^ LENGTH('b))) (y mod (2 ^ LENGTH('b)))) mod (2 ^ LENGTH('b)) = (L x y) mod (2 ^ LENGTH('b))" assumes is_down: "is_down (ucast :: 'a word \ 'b word)" shows "ucast (M a b) = M' (ucast a) (ucast b)" apply (clarsimp simp: word_of_int ucast_def) apply (subst lift_M) apply (subst of_int_uint [symmetric], subst lift_M') apply (subst (1 2) int_word_uint) apply (subst word_of_int) apply (subst word.abs_eq_iff) apply (subst (1 2) bintrunc_mod2p) apply (insert is_down) apply (unfold is_down_def) apply (clarsimp simp: target_size source_size) apply (clarsimp simp: mod_mod_power_int min_def) apply (rule distrib [symmetric]) done lemma ucast_down_add: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) + b) = (ucast a + ucast b :: 'b::len word)" by (rule ucast_distrib [where L="(+)"], (clarsimp simp: uint_word_ariths)+, presburger, simp) lemma ucast_down_minus: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) - b) = (ucast a - ucast b :: 'b::len word)" apply (rule ucast_distrib [where L="(-)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_diff_left_eq mod_diff_right_eq) apply simp done lemma ucast_down_mult: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) * b) = (ucast a * ucast b :: 'b::len word)" apply (rule ucast_distrib [where L="(*)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_mult_eq) apply simp done lemma scast_distrib: fixes M :: "'a::len word \ 'a::len word \ 'a::len word" fixes M' :: "'b::len word \ 'b::len word \ 'b::len word" fixes L :: "int \ int \ int" assumes lift_M: "\x y. uint (M x y) = L (uint x) (uint y) mod 2 ^ LENGTH('a)" assumes lift_M': "\x y. uint (M' x y) = L (uint x) (uint y) mod 2 ^ LENGTH('b)" assumes distrib: "\x y. (L (x mod (2 ^ LENGTH('b))) (y mod (2 ^ LENGTH('b)))) mod (2 ^ LENGTH('b)) = (L x y) mod (2 ^ LENGTH('b))" assumes is_down: "is_down (scast :: 'a word \ 'b word)" shows "scast (M a b) = M' (scast a) (scast b)" apply (subst (1 2 3) down_cast_same [symmetric]) apply (insert is_down) apply (clarsimp simp: is_down_def target_size source_size is_down) apply (rule ucast_distrib [where L=L, OF lift_M lift_M' distrib]) apply (insert is_down) apply (clarsimp simp: is_down_def target_size source_size is_down) done lemma scast_down_add: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) + b) = (scast a + scast b :: 'b::len word)" by (rule scast_distrib [where L="(+)"], (clarsimp simp: uint_word_ariths)+, presburger, simp) lemma scast_down_minus: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) - b) = (scast a - scast b :: 'b::len word)" apply (rule scast_distrib [where L="(-)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_diff_left_eq mod_diff_right_eq) apply simp done lemma scast_down_mult: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) * b) = (scast a * scast b :: 'b::len word)" apply (rule scast_distrib [where L="(*)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_mult_eq) apply simp done lemma scast_ucast_1: "\ is_down (ucast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_def ucast_down_wi) lemma scast_ucast_3: "\ is_down (ucast :: 'a word \ 'c word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_def ucast_down_wi) lemma scast_ucast_4: "\ is_up (ucast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_def ucast_down_wi) lemma scast_scast_b: "\ is_up (scast :: 'a word \ 'b word) \ \ (scast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_def sint_up_scast) lemma ucast_scast_1: "\ is_down (scast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_def ucast_down_wi) lemma ucast_scast_3: "\ is_down (scast :: 'a word \ 'c word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_def ucast_down_wi) lemma ucast_scast_4: "\ is_up (scast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis down_cast_same scast_def sint_up_scast) lemma ucast_ucast_a: "\ is_down (ucast :: 'b word \ 'c word) \ \ (ucast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_def ucast_down_wi) lemma ucast_ucast_b: "\ is_up (ucast :: 'a word \ 'b word) \ \ (ucast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis ucast_up_ucast) lemma scast_scast_a: "\ is_down (scast :: 'b word \ 'c word) \ \ (scast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" apply (clarsimp simp: scast_def) apply (metis down_cast_same is_up_down scast_def ucast_down_wi) done lemma scast_down_wi [OF refl]: "uc = scast \ is_down uc \ uc (word_of_int x) = word_of_int x" by (metis down_cast_same is_up_down ucast_down_wi) lemmas cast_simps = is_down is_up scast_down_add scast_down_minus scast_down_mult ucast_down_add ucast_down_minus ucast_down_mult scast_ucast_1 scast_ucast_3 scast_ucast_4 ucast_scast_1 ucast_scast_3 ucast_scast_4 ucast_ucast_a ucast_ucast_b scast_scast_a scast_scast_b ucast_down_bl ucast_down_wi scast_down_wi ucast_of_nat scast_of_nat uint_up_ucast sint_up_scast up_scast_surj up_ucast_surj lemma smod_mod_positive: "\ 0 \ (a :: int); 0 \ b \ \ a smod b = a mod b" by (clarsimp simp: smod_int_alt_def zsgn_def) lemma nat_mult_power_less_eq: "b > 0 \ (a * b ^ n < (b :: nat) ^ m) = (a < b ^ (m - n))" using mult_less_cancel2[where m = a and k = "b ^ n" and n="b ^ (m - n)"] mult_less_cancel2[where m="a * b ^ (n - m)" and k="b ^ m" and n=1] apply (simp only: power_add[symmetric] nat_minus_add_max) apply (simp only: power_add[symmetric] nat_minus_add_max ac_simps) apply (simp add: max_def split: if_split_asm) done lemma signed_shift_guard_to_word: "\ n < len_of TYPE ('a); n > 0 \ \ (unat (x :: 'a :: len word) * 2 ^ y < 2 ^ n) = (x = 0 \ x < (1 << n >> y))" apply (simp only: nat_mult_power_less_eq) apply (cases "y \ n") apply (simp only: shiftl_shiftr1) apply (subst less_mask_eq) apply (simp add: word_less_nat_alt word_size) apply (rule order_less_le_trans[rotated], rule power_increasing[where n=1]) apply simp apply simp apply simp apply (simp add: nat_mult_power_less_eq word_less_nat_alt word_size) apply auto[1] apply (simp only: shiftl_shiftr2, simp add: unat_eq_0) done lemma sint_ucast_eq_uint: "\ \ is_down (ucast :: ('a::len word \ 'b::len word)) \ \ sint ((ucast :: ('a::len word \ 'b::len word)) x) = uint x" apply (subst sint_eq_uint) apply (clarsimp simp: msb_nth nth_ucast is_down) apply (metis Suc_leI Suc_pred bang_conj_lt len_gt_0) apply (clarsimp simp: uint_up_ucast is_up is_down) done lemma word_less_nowrapI': "(x :: 'a :: len0 word) \ z - k \ k \ z \ 0 < k \ x < x + k" by uint_arith lemma mask_plus_1: "mask n + 1 = 2 ^ n" by (clarsimp simp: mask_def) lemma unat_inj: "inj unat" by (metis eq_iff injI word_le_nat_alt) lemma unat_ucast_upcast: "is_up (ucast :: 'b word \ 'a word) \ unat (ucast x :: ('a::len) word) = unat (x :: ('b::len) word)" unfolding ucast_def unat_def apply (subst int_word_uint) apply (subst mod_pos_pos_trivial) apply simp apply (rule lt2p_lem) apply (clarsimp simp: is_up) apply simp done lemma ucast_mono: "\ (x :: 'b :: len word) < y; y < 2 ^ LENGTH('a) \ \ ucast x < ((ucast y) :: 'a :: len word)" apply (simp add: ucast_nat_def [symmetric]) apply (rule of_nat_mono_maybe) apply (rule unat_less_helper) apply (simp add: Power.of_nat_power) apply (simp add: word_less_nat_alt) done lemma ucast_mono_le: "\x \ y; y < 2 ^ LENGTH('b)\ \ (ucast (x :: 'a :: len word) :: 'b :: len word) \ ucast y" apply (simp add: ucast_nat_def [symmetric]) apply (subst of_nat_mono_maybe_le[symmetric]) apply (rule unat_less_helper) apply (simp add: Power.of_nat_power) apply (rule unat_less_helper) apply (erule le_less_trans) apply (simp add: Power.of_nat_power) apply (simp add: word_le_nat_alt) done lemma ucast_mono_le': "\ unat y < 2 ^ LENGTH('b); LENGTH('b::len) < LENGTH('a::len); x \ y \ \ UCAST('a \ 'b) x \ UCAST('a \ 'b) y" by (auto simp: word_less_nat_alt intro: ucast_mono_le) lemma zero_sle_ucast_up: "\ is_down (ucast :: 'a word \ 'b signed word) \ (0 <=s ((ucast (b::('a::len) word)) :: ('b::len) signed word))" apply (subgoal_tac "\ msb (ucast b :: 'b signed word)") apply (clarsimp simp: word_sle_msb_le) apply (clarsimp simp: is_down not_le msb_nth nth_ucast) apply (subst (asm) bang_conj_lt [symmetric]) apply clarsimp apply arith done lemma word_le_ucast_sless: "\ x \ y; y \ -1; LENGTH('a) < LENGTH('b) \ \ UCAST (('a :: len) \ ('b :: len) signed) x msb (ucast x :: ('a::len) word) = msb (x :: ('b::len) word)" apply (clarsimp simp: word_msb_alt) apply (subst ucast_down_drop [where n=0]) apply (clarsimp simp: source_size_def target_size_def word_size) apply clarsimp done lemma msb_big: "msb (a :: ('a::len) word) = (a \ 2 ^ (LENGTH('a) - Suc 0))" apply (rule iffI) apply (clarsimp simp: msb_nth) apply (drule bang_is_le) apply simp apply (rule ccontr) apply (subgoal_tac "a = a && mask (LENGTH('a) - Suc 0)") apply (cut_tac and_mask_less' [where w=a and n="LENGTH('a) - Suc 0"]) apply (clarsimp simp: word_not_le [symmetric]) apply clarsimp apply (rule sym, subst and_mask_eq_iff_shiftr_0) apply (clarsimp simp: msb_shift) done lemma zero_sle_ucast: "(0 <=s ((ucast (b::('a::len) word)) :: ('a::len) signed word)) = (uint b < 2 ^ (len_of (TYPE('a)) - 1))" apply (case_tac "msb b") apply (clarsimp simp: word_sle_msb_le not_less msb_ucast_eq del: notI) apply (clarsimp simp: msb_big word_le_def uint_2p_alt) apply (clarsimp simp: word_sle_msb_le not_less msb_ucast_eq del: notI) apply (clarsimp simp: msb_big word_le_def uint_2p_alt) done (* to_bool / from_bool. *) definition from_bool :: "bool \ 'a::len word" where "from_bool b \ case b of True \ of_nat 1 | False \ of_nat 0" lemma from_bool_0: "(from_bool x = 0) = (\ x)" by (simp add: from_bool_def split: bool.split) definition to_bool :: "'a::len word \ bool" where "to_bool \ (\) 0" lemma to_bool_and_1: "to_bool (x && 1) = (x !! 0)" apply (simp add: to_bool_def) apply (rule iffI) apply (rule classical, erule notE, rule word_eqI) apply clarsimp apply (case_tac n, simp_all)[1] apply (rule notI, drule word_eqD[where x=0]) apply simp done lemma to_bool_from_bool [simp]: "to_bool (from_bool r) = r" unfolding from_bool_def to_bool_def by (simp split: bool.splits) lemma from_bool_neq_0 [simp]: "(from_bool b \ 0) = b" by (simp add: from_bool_def split: bool.splits) lemma from_bool_mask_simp [simp]: "(from_bool r :: 'a::len word) && 1 = from_bool r" unfolding from_bool_def by (clarsimp split: bool.splits) lemma from_bool_1 [simp]: "(from_bool P = 1) = P" by (simp add: from_bool_def split: bool.splits) lemma ge_0_from_bool [simp]: "(0 < from_bool P) = P" by (simp add: from_bool_def split: bool.splits) lemma limited_and_from_bool: "limited_and (from_bool b) 1" by (simp add: from_bool_def limited_and_def split: bool.split) lemma to_bool_1 [simp]: "to_bool 1" by (simp add: to_bool_def) lemma to_bool_0 [simp]: "\to_bool 0" by (simp add: to_bool_def) lemma from_bool_eq_if: "(from_bool Q = (if P then 1 else 0)) = (P = Q)" by (simp add: case_bool_If from_bool_def split: if_split) lemma to_bool_eq_0: "(\ to_bool x) = (x = 0)" by (simp add: to_bool_def) lemma to_bool_neq_0: "(to_bool x) = (x \ 0)" by (simp add: to_bool_def) lemma from_bool_all_helper: "(\bool. from_bool bool = val \ P bool) = ((\bool. from_bool bool = val) \ P (val \ 0))" by (auto simp: from_bool_0) lemma from_bool_to_bool_iff: "w = from_bool b \ to_bool w = b \ (w = 0 \ w = 1)" by (cases b) (auto simp: from_bool_def to_bool_def) lemma word_rsplit_upt: "\ size x = LENGTH('a :: len) * n; n \ 0 \ \ word_rsplit x = map (\i. ucast (x >> i * len_of TYPE ('a)) :: 'a word) (rev [0 ..< n])" apply (subgoal_tac "length (word_rsplit x :: 'a word list) = n") apply (rule nth_equalityI, simp) apply (intro allI word_eqI impI) apply (simp add: test_bit_rsplit_alt word_size) apply (simp add: nth_ucast nth_shiftr nth_rev field_simps) apply (simp add: length_word_rsplit_exp_size) apply (metis mult.commute given_quot_alt word_size word_size_gt_0) done lemma aligned_shift: "\x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \ LENGTH('a)\ \ x + y >> n = y >> n" apply (subst word_plus_and_or_coroll) apply (rule word_eqI) apply (clarsimp simp: is_aligned_nth) apply (drule(1) nth_bounded) apply simp apply simp apply (rule word_eqI) apply (simp add: nth_shiftr) apply safe apply (drule(1) nth_bounded) apply simp+ done lemma aligned_shift': "\x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \ LENGTH('a)\ \ y + x >> n = y >> n" apply (subst word_plus_and_or_coroll) apply (rule word_eqI) apply (clarsimp simp: is_aligned_nth) apply (drule(1) nth_bounded) apply simp apply simp apply (rule word_eqI) apply (simp add: nth_shiftr) apply safe apply (drule(1) nth_bounded) apply simp+ done lemma neg_mask_add_mask: "((x:: 'a :: len word) && ~~ mask n) + (2 ^ n - 1) = x || mask n" apply (simp add:mask_2pm1[symmetric]) apply (rule word_eqI[rule_format]) apply (rule iffI) apply (clarsimp simp:word_size not_less) apply (cut_tac w = "((x && ~~ mask n) + mask n)" and m = n and n = "na - n" in nth_shiftr[symmetric]) apply clarsimp apply (subst (asm) aligned_shift') apply (simp add:mask_lt_2pn nth_shiftr is_aligned_neg_mask word_size)+ apply (case_tac "na p" apply (rule word_le_minus_cancel[where x = "p && ~~ mask n"]) apply (clarsimp simp: subtract_mask) using word_and_le1[where a = "mask n" and y = p] apply (clarsimp simp: mask_def word_le_less_eq) apply (rule is_aligned_no_overflow'[folded mask_2pm1]) apply (clarsimp simp: is_aligned_neg_mask) done lemma word_neg_and_le: "ptr \ (ptr && ~~ mask n) + (2 ^ n - 1)" by (simp add: and_neg_mask_plus_mask_mono mask_2pm1[symmetric]) lemma aligned_less_plus_1: "\ is_aligned x n; n > 0 \ \ x < x + 1" apply (rule plus_one_helper2) apply (rule order_refl) apply (clarsimp simp: field_simps) apply (drule arg_cong[where f="\x. x - 1"]) apply (clarsimp simp: is_aligned_mask) apply (drule word_eqD[where x=0]) apply simp done lemma aligned_add_offset_less: "\is_aligned x n; is_aligned y n; x < y; z < 2 ^ n\ \ x + z < y" apply (cases "y = 0") apply simp apply (erule is_aligned_get_word_bits[where p=y], simp_all) apply (cases "z = 0", simp_all) apply (drule(2) aligned_at_least_t2n_diff[rotated -1]) apply (drule plus_one_helper2) apply (rule less_is_non_zero_p1) apply (rule aligned_less_plus_1) apply (erule aligned_sub_aligned[OF _ _ order_refl], simp_all add: is_aligned_triv)[1] apply (cases n, simp_all)[1] apply (simp only: trans[OF diff_add_eq diff_diff_eq2[symmetric]]) apply (drule word_less_add_right) apply (rule ccontr, simp add: linorder_not_le) apply (drule aligned_small_is_0, erule order_less_trans) apply (clarsimp simp: power_overflow) apply simp apply (erule order_le_less_trans[rotated], rule word_plus_mono_right) apply (erule minus_one_helper3) apply (simp add: is_aligned_no_wrap' is_aligned_no_overflow field_simps) done lemma is_aligned_add_helper: "\ is_aligned p n; d < 2 ^ n \ \ (p + d && mask n = d) \ (p + d && (~~ mask n) = p)" apply (subst(asm) is_aligned_mask) apply (drule less_mask_eq) apply (rule context_conjI) apply (subst word_plus_and_or_coroll) apply (rule word_eqI) apply (drule_tac x=na in word_eqD)+ apply (simp add: word_size) apply blast apply (rule word_eqI) apply (drule_tac x=na in word_eqD)+ apply (simp add: word_ops_nth_size word_size) apply blast apply (insert word_plus_and_or_coroll2[where x="p + d" and w="mask n"]) apply simp done lemma is_aligned_sub_helper: "\ is_aligned (p - d) n; d < 2 ^ n \ \ (p && mask n = d) \ (p && (~~ mask n) = p - d)" by (drule(1) is_aligned_add_helper, simp) lemma mask_twice: "(x && mask n) && mask m = x && mask (min m n)" apply (rule word_eqI) apply (simp add: word_size conj_comms) done lemma is_aligned_after_mask: "\is_aligned k m;m\ n\ \ is_aligned (k && mask n) m" by (rule is_aligned_andI1) lemma and_mask_plus: "\is_aligned ptr m; m \ n; a < 2 ^ m\ \ ptr + a && mask n = (ptr && mask n) + a" apply (rule mask_eqI[where n = m]) apply (simp add:mask_twice min_def) apply (simp add:is_aligned_add_helper) apply (subst is_aligned_add_helper[THEN conjunct1]) apply (erule is_aligned_after_mask) apply simp apply simp apply simp apply (subgoal_tac "(ptr + a && mask n) && ~~ mask m = (ptr + a && ~~ mask m ) && mask n") apply (simp add:is_aligned_add_helper) apply (subst is_aligned_add_helper[THEN conjunct2]) apply (simp add:is_aligned_after_mask) apply simp apply simp apply (simp add:word_bw_comms word_bw_lcs) done lemma le_step_down_word:"\(i::('a::len) word) \ n; i = n \ P; i \ n - 1 \ P\ \ P" by unat_arith lemma le_step_down_word_2: fixes x :: "'a::len word" shows "\x \ y; x \ y\ \ x \ y - 1" by (subst (asm) word_le_less_eq, clarsimp, simp add: minus_one_helper3) lemma NOT_mask_AND_mask[simp]: "(w && mask n) && ~~ mask n = 0" apply (clarsimp simp:mask_def) by (metis word_bool_alg.conj_cancel_right word_bool_alg.conj_zero_right word_bw_comms(1) word_bw_lcs(1)) lemma and_and_not[simp]:"(a && b) && ~~ b = 0" apply (subst word_bw_assocs(1)) apply clarsimp done lemma mask_shift_and_negate[simp]:"(w && mask n << m) && ~~ (mask n << m) = 0" apply (clarsimp simp:mask_def) by (metis (erased, hide_lams) mask_eq_x_eq_0 shiftl_over_and_dist word_bool_alg.conj_absorb word_bw_assocs(1)) lemma le_step_down_nat:"\(i::nat) \ n; i = n \ P; i \ n - 1 \ P\ \ P" by arith lemma le_step_down_int:"\(i::int) \ n; i = n \ P; i \ n - 1 \ P\ \ P" by arith lemma mask_1[simp]: "(\x. mask x = 1)" apply (rule_tac x=1 in exI) apply (simp add:mask_def) done lemma not_switch:"~~ a = x \ a = ~~ x" by auto (* The seL4 bitfield generator produces functions containing mask and shift operations, such that * invoking two of them consecutively can produce something like the following. *) lemma bitfield_op_twice: "(x && ~~ (mask n << m) || ((y && mask n) << m)) && ~~ (mask n << m) = x && ~~ (mask n << m)" by (induct n arbitrary: m) (auto simp: word_ao_dist) lemma bitfield_op_twice'': "\~~ a = b << c; \x. b = mask x\ \ (x && a || (y && b << c)) && a = x && a" apply clarsimp apply (cut_tac n=xa and m=c and x=x and y=y in bitfield_op_twice) apply (clarsimp simp:mask_def) apply (drule not_switch) apply clarsimp done lemma bit_twiddle_min: "(y::'a::len word) xor (((x::'a::len word) xor y) && (if x < y then -1 else 0)) = min x y" by (metis (mono_tags) min_def word_bitwise_m1_simps(2) word_bool_alg.xor_left_commute word_bool_alg.xor_self word_bool_alg.xor_zero_right word_bw_comms(1) word_le_less_eq word_log_esimps(7)) lemma bit_twiddle_max: "(x::'a::len word) xor (((x::'a::len word) xor y) && (if x < y then -1 else 0)) = max x y" by (metis (mono_tags) max_def word_bitwise_m1_simps(2) word_bool_alg.xor_left_self word_bool_alg.xor_zero_right word_bw_comms(1) word_le_less_eq word_log_esimps(7)) lemma swap_with_xor: "\(x::'a::len word) = a xor b; y = b xor x; z = x xor y\ \ z = b \ y = a" by (metis word_bool_alg.xor_assoc word_bool_alg.xor_commute word_bool_alg.xor_self word_bool_alg.xor_zero_right) -lemma scast_nop_1: +lemma scast_nop1: "((scast ((of_int x)::('a::len) word))::'a sword) = of_int x" apply (clarsimp simp:scast_def word_of_int) by (metis len_signed sint_sbintrunc' word_sint.Rep_inverse) -lemma scast_nop_2: +lemma scast_nop2: "((scast ((of_int x)::('a::len) sword))::'a word) = of_int x" apply (clarsimp simp:scast_def word_of_int) by (metis len_signed sint_sbintrunc' word_sint.Rep_inverse) -lemmas scast_nop[simp] = scast_nop_1 scast_nop_2 scast_id +lemmas scast_nop[simp] = scast_nop1 scast_nop2 scast_id lemma le_mask_imp_and_mask: "(x::'a::len word) \ mask n \ x && mask n = x" by (metis and_mask_eq_iff_le_mask) lemma or_not_mask_nop: "((x::'a::len word) || ~~ mask n) && mask n = x && mask n" by (metis word_and_not word_ao_dist2 word_bw_comms(1) word_log_esimps(3)) lemma mask_subsume: "\n \ m\ \ ((x::'a::len word) || y && mask n) && ~~ mask m = x && ~~ mask m" apply (subst word_ao_dist) apply (subgoal_tac "(y && mask n) && ~~ mask m = 0") apply simp by (metis (no_types, hide_lams) is_aligned_mask is_aligned_weaken word_and_not word_bool_alg.conj_zero_right word_bw_comms(1) word_bw_lcs(1)) lemma mask_twice2: "n \ m \ ((x::'a::len word) && mask m) && mask n = x && mask n" by (metis mask_twice min_def) lemma uint_2_id: "LENGTH('a) \ 2 \ uint (2::('a::len) word) = 2" apply clarsimp apply (subgoal_tac "2 \ uints (LENGTH('a))") apply (subst (asm) Word.word_ubin.set_iff_norm) apply simp apply (subst word_uint.set_iff_norm) apply clarsimp apply (rule int_mod_eq') apply simp apply (rule pow_2_gt) apply simp done lemma bintrunc_id: "\m \ of_nat n; 0 < m\ \ bintrunc n m = m" by (simp add: bintrunc_mod2p le_less_trans) lemma shiftr1_unfold: "shiftr1 x = x >> 1" by (metis One_nat_def comp_apply funpow.simps(1) funpow.simps(2) id_apply shiftr_def) lemma shiftr1_is_div_2: "(x::('a::len) word) >> 1 = x div 2" apply (case_tac "LENGTH('a) = 1") apply simp apply (subgoal_tac "x = 0 \ x = 1") apply (erule disjE) apply (clarsimp simp:word_div_def)+ apply (metis One_nat_def less_irrefl_nat sint_1_cases) apply clarsimp apply (subst word_div_def) apply clarsimp apply (subst bintrunc_id) apply (subgoal_tac "2 \ LENGTH('a)") apply simp apply (metis (no_types) le_0_eq le_SucE lens_not_0(2) not_less_eq_eq numeral_2_eq_2) apply simp apply (subst bin_rest_def[symmetric]) apply (subst shiftr1_def[symmetric]) apply (clarsimp simp:shiftr1_unfold) done lemma shiftl1_is_mult: "(x << 1) = (x :: 'a::len word) * 2" by (metis One_nat_def mult_2 mult_2_right one_add_one power_0 power_Suc shiftl_t2n) lemma div_of_0_id[simp]:"(0::('a::len) word) div n = 0" by (simp add: word_div_def) lemma degenerate_word:"LENGTH('a) = 1 \ (x::('a::len) word) = 0 \ x = 1" by (metis One_nat_def less_irrefl_nat sint_1_cases) lemma div_by_0_word:"(x::('a::len) word) div 0 = 0" by (metis div_0 div_by_0 unat_0 word_arith_nat_defs(6) word_div_1) lemma div_less_dividend_word:"\x \ 0; n \ 1\ \ (x::('a::len) word) div n < x" apply (case_tac "n = 0") apply clarsimp apply (subst div_by_0_word) apply (simp add:word_neq_0_conv) apply (subst word_arith_nat_div) apply (rule word_of_nat_less) apply (rule div_less_dividend) apply (metis (poly_guards_query) One_nat_def less_one nat_neq_iff unat_eq_1(2) unat_eq_zero) apply (simp add:unat_gt_0) done lemma shiftr1_lt:"x \ 0 \ (x::('a::len) word) >> 1 < x" apply (subst shiftr1_is_div_2) apply (rule div_less_dividend_word) apply simp+ done lemma word_less_div: fixes x :: "('a::len) word" and y :: "('a::len) word" shows "x div y = 0 \ y = 0 \ x < y" apply (case_tac "y = 0", clarsimp+) by (metis One_nat_def Suc_le_mono le0 le_div_geq not_less unat_0 unat_div unat_gt_0 word_less_nat_alt zero_less_one) lemma not_degenerate_imp_2_neq_0:"LENGTH('a) > 1 \ (2::('a::len) word) \ 0" by (metis numerals(1) power_not_zero power_zero_numeral) lemma shiftr1_0_or_1:"(x::('a::len) word) >> 1 = 0 \ x = 0 \ x = 1" apply (subst (asm) shiftr1_is_div_2) apply (drule word_less_div) apply (case_tac "LENGTH('a) = 1") apply (simp add:degenerate_word) apply (erule disjE) apply (subgoal_tac "(2::'a word) \ 0") apply simp apply (rule not_degenerate_imp_2_neq_0) apply (subgoal_tac "LENGTH('a) \ 0") apply arith apply simp apply (rule x_less_2_0_1', simp+) done lemma word_overflow:"(x::('a::len) word) + 1 > x \ x + 1 = 0" apply clarsimp by (metis diff_0 eq_diff_eq less_x_plus_1 max_word_max plus_1_less) lemma word_overflow_unat:"unat ((x::('a::len) word) + 1) = unat x + 1 \ x + 1 = 0" by (metis Suc_eq_plus1 add.commute unatSuc) lemma even_word_imp_odd_next:"even (unat (x::('a::len) word)) \ x + 1 = 0 \ odd (unat (x + 1))" apply (cut_tac x=x in word_overflow_unat) apply clarsimp done lemma odd_word_imp_even_next:"odd (unat (x::('a::len) word)) \ x + 1 = 0 \ even (unat (x + 1))" apply (cut_tac x=x in word_overflow_unat) apply clarsimp done lemma overflow_imp_lsb:"(x::('a::len) word) + 1 = 0 \ x !! 0" by (metis add.commute add_left_cancel max_word_max not_less word_and_1 word_bool_alg.conj_one_right word_bw_comms(1) word_overflow zero_neq_one) lemma word_lsb_nat:"lsb w = (unat w mod 2 = 1)" unfolding word_lsb_def bin_last_def by (metis (no_types, hide_lams) nat_mod_distrib nat_numeral not_mod_2_eq_1_eq_0 numeral_One uint_eq_0 uint_nonnegative unat_0 unat_def zero_le_numeral) lemma odd_iff_lsb:"odd (unat (x::('a::len) word)) = x !! 0" apply (simp add:even_iff_mod_2_eq_zero) apply (subst word_lsb_nat[unfolded One_nat_def, symmetric]) apply (rule word_lsb_alt) done lemma of_nat_neq_iff_word: "x mod 2 ^ LENGTH('a) \ y mod 2 ^ LENGTH('a) \ (((of_nat x)::('a::len) word) \ of_nat y) = (x \ y)" apply (rule iffI) apply (case_tac "x = y") apply (subst (asm) of_nat_eq_iff[symmetric]) apply simp+ apply (case_tac "((of_nat x)::('a::len) word) = of_nat y") apply (subst (asm) word_unat.norm_eq_iff[symmetric]) apply simp+ done lemma shiftr1_irrelevant_lsb:"(x::('a::len) word) !! 0 \ x >> 1 = (x + 1) >> 1" apply (case_tac "LENGTH('a) = 1") apply clarsimp apply (drule_tac x=x in degenerate_word[unfolded One_nat_def]) apply (erule disjE) apply clarsimp+ apply (subst (asm) shiftr1_is_div_2[unfolded One_nat_def])+ apply (subst (asm) word_arith_nat_div)+ apply clarsimp apply (subst (asm) bintrunc_id) apply (subgoal_tac "LENGTH('a) > 0") apply linarith apply clarsimp+ apply (subst (asm) bintrunc_id) apply (subgoal_tac "LENGTH('a) > 0") apply linarith apply clarsimp+ apply (case_tac "x + 1 = 0") apply (clarsimp simp:overflow_imp_lsb) apply (cut_tac x=x in word_overflow_unat) apply clarsimp apply (case_tac "even (unat x)") apply (subgoal_tac "unat x div 2 = Suc (unat x) div 2") apply metis apply (subst numeral_2_eq_2)+ apply simp apply (simp add:odd_iff_lsb) done lemma shiftr1_0_imp_only_lsb:"((x::('a::len) word) + 1) >> 1 = 0 \ x = 0 \ x + 1 = 0" by (metis One_nat_def shiftr1_0_or_1 word_less_1 word_overflow) lemma shiftr1_irrelevant_lsb':"\((x::('a::len) word) !! 0) \ x >> 1 = (x + 1) >> 1" by (metis shiftr1_irrelevant_lsb) lemma lsb_this_or_next:"\(((x::('a::len) word) + 1) !! 0) \ x !! 0" by (metis (poly_guards_query) even_word_imp_odd_next odd_iff_lsb overflow_imp_lsb) (* Perhaps this one should be a simp lemma, but it seems a little dangerous. *) lemma cast_chunk_assemble_id: "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ (((ucast ((ucast (x::'b word))::'a word))::'b word) || (((ucast ((ucast (x >> n))::'a word))::'b word) << n)) = x" apply (subgoal_tac "((ucast ((ucast (x >> n))::'a word))::'b word) = x >> n") apply clarsimp apply (subst and_not_mask[symmetric]) apply (subst ucast_ucast_mask) apply (subst word_ao_dist2[symmetric]) apply clarsimp apply (rule ucast_ucast_len) apply (rule shiftr_less_t2n') apply (subst and_mask_eq_iff_le_mask) apply (clarsimp simp:mask_def) apply (metis max_word_eq max_word_max mult_2_right) apply (metis add_diff_cancel_left' diff_less len_gt_0 mult_2_right) done lemma cast_chunk_scast_assemble_id: "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ (((ucast ((scast (x::'b word))::'a word))::'b word) || (((ucast ((scast (x >> n))::'a word))::'b word) << n)) = x" apply (subgoal_tac "((scast x)::'a word) = ((ucast x)::'a word)") apply (subgoal_tac "((scast (x >> n))::'a word) = ((ucast (x >> n))::'a word)") apply (simp add:cast_chunk_assemble_id) apply (subst down_cast_same[symmetric], subst is_down, arith, simp)+ done lemma mask_or_not_mask: "x && mask n || x && ~~ mask n = x" apply (subst word_oa_dist, simp) apply (subst word_oa_dist2, simp) done lemma is_aligned_add_not_aligned: "\is_aligned (p::'a::len word) n; \ is_aligned (q::'a::len word) n\ \ \ is_aligned (p + q) n" by (metis is_aligned_addD1) lemma word_gr0_conv_Suc: "(m::'a::len word) > 0 \ \n. m = n + 1" by (metis add.commute add_minus_cancel) lemma neg_mask_add_aligned: "\ is_aligned p n; q < 2 ^ n \ \ (p + q) && ~~ mask n = p && ~~ mask n" by (metis is_aligned_add_helper is_aligned_neg_mask_eq) lemma word_sless_sint_le:"x sint x \ sint y - 1" by (metis word_sless_alt zle_diff1_eq) lemma upper_trivial: fixes x :: "'a::len word" shows "x \ 2 ^ LENGTH('a) - 1 \ x < 2 ^ LENGTH('a) - 1" by (cut_tac n=x and 'a='a in max_word_max, clarsimp simp:max_word_def, simp add: less_le) lemma constraint_expand: fixes x :: "'a::len word" shows "x \ {y. lower \ y \ y \ upper} = (lower \ x \ x \ upper)" by (rule mem_Collect_eq) lemma card_map_elide: "card ((of_nat :: nat \ 'a::len word) ` {0.. CARD('a::len word)" proof - let ?of_nat = "of_nat :: nat \ 'a word" from word_unat.Abs_inj_on have "inj_on ?of_nat {i. i < CARD('a word)}" by (simp add: unats_def card_word) moreover have "{0.. {i. i < CARD('a word)}" using that by auto ultimately have "inj_on ?of_nat {0.. CARD('a::len word) \ card ((of_nat::nat \ 'a::len word) ` {0.. ucast (max_word::'b::len word)" shows "ucast ((ucast x)::'b word) = x" proof - { assume a1: "x \ word_of_int (uint (word_of_int (2 ^ len_of (TYPE('b)) - 1)::'b word))" have f2: "((\i ia. (0::int) \ i \ \ 0 \ i + - 1 * ia \ i mod ia \ i) \ \ (0::int) \ - 1 + 2 ^ LENGTH('b) \ (0::int) \ - 1 + 2 ^ LENGTH('b) + - 1 * 2 ^ LENGTH('b) \ (- (1::int) + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b) = - 1 + 2 ^ LENGTH('b)) = ((\i ia. (0::int) \ i \ \ 0 \ i + - 1 * ia \ i mod ia \ i) \ \ (1::int) \ 2 ^ LENGTH('b) \ 2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)) = 1)" by force have f3: "\i ia. \ (0::int) \ i \ 0 \ i + - 1 * ia \ i mod ia = i" using mod_pos_pos_trivial by force have "(1::int) \ 2 ^ LENGTH('b)" by simp then have "2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ len_of TYPE ('b)) = 1" using f3 f2 by blast then have f4: "- (1::int) + 2 ^ LENGTH('b) = (- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)" by linarith have f5: "x \ word_of_int (uint (word_of_int (- 1 + 2 ^ LENGTH('b))::'b word))" using a1 by force have f6: "2 ^ len_of (TYPE('b)::'b itself) + - (1::int) = - 1 + 2 ^ len_of (TYPE('b))" by force have f7: "- (1::int) * 1 = - 1" by auto have "\x0 x1. (x1::int) - x0 = x1 + - 1 * x0" by force then have "x \ 2 ^ LENGTH('b) - 1" using f7 f6 f5 f4 by (metis uint_word_of_int wi_homs(2) word_arith_wis(8) word_of_int_2p) } with assms show ?thesis unfolding ucast_def apply (subst word_ubin.eq_norm) apply (subst and_mask_bintr[symmetric]) apply (subst and_mask_eq_iff_le_mask) apply (clarsimp simp: max_word_def mask_def) done qed lemma remdups_enum_upto: fixes s::"'a::len word" shows "remdups [s .e. e] = [s .e. e]" by simp lemma card_enum_upto: fixes s::"'a::len word" shows "card (set [s .e. e]) = Suc (unat e) - unat s" by (subst List.card_set) (simp add: remdups_enum_upto) lemma unat_mask: "unat (mask n :: 'a :: len word) = 2 ^ (min n (LENGTH('a))) - 1" apply (subst min.commute) apply (simp add: mask_def not_less min_def split: if_split_asm) apply (intro conjI impI) apply (simp add: unat_sub_if_size) apply (simp add: power_overflow word_size) apply (simp add: unat_sub_if_size) done lemma word_shiftr_lt: fixes w :: "'a::len word" shows "unat (w >> n) < (2 ^ (len_of(TYPE('a)) - n))" apply (subst shiftr_div_2n') by (metis nat_mod_lem nat_zero_less_power_iff power_mod_div word_unat.Rep_inverse word_unat.eq_norm zero_less_numeral) lemma complement_nth_w2p: shows "n' < len_of(TYPE('a)) \ (~~ (2 ^ n :: ('a::len word))) !! n' = (n' \ n)" by (fastforce simp: word_ops_nth_size word_size nth_w2p) lemma word_unat_and_lt: "unat x < n \ unat y < n \ unat (x && y) < n" by (meson le_less_trans word_and_le1 word_and_le2 word_le_nat_alt) lemma word_unat_mask_lt: "m \ size w \ unat ((w::'a::len word) && mask m) < 2 ^ m" by (rule word_unat_and_lt) (simp add: unat_mask word_size) lemma unat_shiftr_less_t2n: fixes x :: "'a :: len word" shows "unat x < 2 ^ (n + m) \ unat (x >> n) < 2 ^ m" by (simp add: shiftr_div_2n' power_add mult.commute td_gal_lt) lemma le_or_mask: "w \ w' \ w || mask x \ w' || mask x" by (metis neg_mask_add_mask add.commute le_word_or1 mask_2pm1 neg_mask_mono_le word_plus_mono_left) lemma le_shiftr1': "\ shiftr1 u \ shiftr1 v ; shiftr1 u \ shiftr1 v \ \ u \ v" apply (unfold word_le_def shiftr1_def word_ubin.eq_norm) apply (unfold bin_rest_trunc_i trans [OF bintrunc_bintrunc_l word_ubin.norm_Rep, unfolded word_ubin.norm_Rep, OF order_refl [THEN le_SucI]]) apply (case_tac "uint u" rule: bin_exhaust) apply (rename_tac bs bit) apply (case_tac "uint v" rule: bin_exhaust) apply (rename_tac bs' bit') apply (case_tac "bit") apply (case_tac "bit'", auto simp: less_eq_int_code le_Bits intro: basic_trans_rules)[1] apply (case_tac bit') apply (simp add: le_Bits less_eq_int_code) apply (auto simp: le_Bits less_eq_int_code) done lemma le_shiftr': "\ u >> n \ v >> n ; u >> n \ v >> n \ \ (u::'a::len0 word) \ v" apply (induct n, simp) apply (unfold shiftr_def) apply (case_tac "(shiftr1 ^^ n) u = (shiftr1 ^^ n) v") apply simp apply (fastforce dest: le_shiftr1') done lemma word_log2_nth_same: "w \ 0 \ w !! word_log2 w" unfolding word_log2_def using nth_length_takeWhile[where P=Not and xs="to_bl w"] apply (simp add: word_clz_def word_size to_bl_nth) apply (fastforce simp: linorder_not_less eq_zero_set_bl dest: takeWhile_take_has_property) done lemma word_log2_nth_not_set: "\ word_log2 w < i ; i < size w \ \ \ w !! i" unfolding word_log2_def word_clz_def using takeWhile_take_has_property_nth[where P=Not and xs="to_bl w" and n="size w - Suc i"] by (fastforce simp add: to_bl_nth word_size) lemma word_log2_highest: assumes a: "w !! i" shows "i \ word_log2 w" proof - from a have "i < size w" by - (rule test_bit_size) with a show ?thesis by - (rule ccontr, simp add: word_log2_nth_not_set) qed lemma word_log2_max: "word_log2 w < size w" unfolding word_log2_def word_clz_def by simp lemma word_clz_0[simp]: "word_clz (0::'a::len word) \ len_of (TYPE('a))" unfolding word_clz_def by (simp add: takeWhile_replicate) lemma word_clz_minus_one[simp]: "word_clz (-1::'a::len word) \ 0" unfolding word_clz_def by (simp add: max_word_bl[simplified max_word_minus] takeWhile_replicate) lemma word_add_no_overflow:"(x::'a::len word) < max_word \ x < x + 1" using less_x_plus_1 order_less_le by blast lemma lt_plus_1_le_word: fixes x :: "'a::len word" assumes bound:"n < unat (maxBound::'a word)" shows "x < 1 + of_nat n = (x \ of_nat n)" apply (subst le_less) apply (rule iffI) apply clarsimp apply (subst (asm) add.commute) apply (subst (asm) less_x_plus_1) apply (cut_tac ?'a='a and y=n and x="unat (maxBound::'a word) - 1" in of_nat_mono_maybe') apply clarsimp apply (simp add: less_imp_diff_less) apply (insert bound[unfolded maxBound_word, simplified]) using order_less_trans apply blast apply (metis max_word_minus word_not_simps(3) word_of_nat_less) apply clarsimp apply (erule disjE) apply (subgoal_tac "x < of_nat (1 + n)") prefer 2 apply (cut_tac ?'a='a and y="unat x" and x="1 + n" in of_nat_mono_maybe) apply clarsimp apply (simp add: less_trans_Suc) apply (simp add: less_Suc_eq unat_less_helper) apply clarsimp apply clarsimp apply clarsimp apply (cut_tac y="(of_nat n)::'a word" and x="(of_nat n)::'a word" in less_x_plus_1) apply (cut_tac ?'a='a and y=n and x="unat (maxBound::'a word) - 1" in of_nat_mono_maybe') apply clarsimp apply (simp add: less_imp_diff_less) apply (insert bound[unfolded maxBound_word, simplified]) using order_less_trans apply blast apply (metis max_word_minus word_not_simps(3) word_of_nat_less) apply (clarsimp simp: add.commute) done lemma unat_ucast_up_simp: fixes x :: "'a::len word" assumes uc :"LENGTH('a) \ LENGTH('b)" shows "unat (ucast x :: 'b::len word) = unat x" unfolding ucast_def unat_def apply (subst int_word_uint) apply (subst mod_pos_pos_trivial) apply simp apply (rule lt2p_lem) apply (simp add: uc) apply simp done lemma unat_ucast_less_no_overflow: "\n < 2 ^ LENGTH('a); unat f < n\ \ (f::('a::len) word) < of_nat n" apply (erule order_le_less_trans[OF _ of_nat_mono_maybe,rotated]) apply assumption apply simp done lemma unat_ucast_less_no_overflow_simp: "n < 2 ^ LENGTH('a) \ (unat f < n) = ((f::('a::len) word) < of_nat n)" apply (rule iffI) apply (simp add:unat_ucast_less_no_overflow) apply (simp add:unat_less_helper) done lemma unat_ucast_no_overflow_le: assumes no_overflow : "unat b < (2 :: nat) ^ LENGTH('a)" and upward_cast: "LENGTH('a) < LENGTH('b)" shows "(ucast (f::('a::len) word) < (b :: 'b :: len word)) = (unat f < unat b)" proof - have LR: "ucast f < b \ unat f < unat b" apply (rule unat_less_helper) apply (simp add:ucast_nat_def) apply (rule_tac 'b1 = 'b in ucast_less_ucast[THEN iffD1]) apply (rule upward_cast) apply (simp add: ucast_ucast_mask less_mask_eq word_less_nat_alt unat_power_lower[OF upward_cast] no_overflow) done have RL: "unat f < unat b \ ucast f < b" proof- assume ineq: "unat f < unat b" have ucast_rewrite: "ucast (f::('a::len) word) < ((ucast (ucast b ::('a::len) word)) :: 'b :: len word)" apply (simp add: ucast_less_ucast upward_cast) apply (simp add: ucast_nat_def[symmetric]) apply (rule unat_ucast_less_no_overflow[OF no_overflow ineq]) done then show ?thesis apply (rule order_less_le_trans) apply (simp add:ucast_ucast_mask word_and_le2) done qed then show ?thesis by (simp add:RL LR iffI) qed (* casting a long word to a shorter word and casting back to the long word is equal to the original long word -- if the word is small enough. 'l is the longer word. 's is the shorter word. *) lemma bl_cast_long_short_long_ingoreLeadingZero_generic: "length (dropWhile Not (to_bl w)) \ LENGTH('s) \ LENGTH('s) \ LENGTH('l) \ (of_bl:: bool list \ 'l::len word) (to_bl ((of_bl:: bool list \ 's::len word) (to_bl w))) = w" apply(rule word_uint_eqI) apply(subst uint_of_bl_is_bl_to_bin) apply(simp; fail) apply(subst to_bl_bin) apply(subst uint_of_bl_is_bl_to_bin_drop) apply blast apply(simp) done (* Casting between longer and shorter word. 'l is the longer word. 's is the shorter word. For example: 'l::len word is 128 word (full ipv6 address) 's::len word is 16 word (address piece of ipv6 address in colon-text-representation) *) corollary ucast_short_ucast_long_ingoreLeadingZero: "length (dropWhile Not (to_bl w)) \ LENGTH('s) \ LENGTH('s) \ LENGTH('l) \ (ucast:: 's::len word \ 'l::len word) ((ucast:: 'l::len word \ 's::len word) w) = w" apply(subst Word.ucast_bl)+ apply(rule bl_cast_long_short_long_ingoreLeadingZero_generic) apply(simp_all) done lemma length_drop_mask: fixes w::"'a::len word" shows "length (dropWhile Not (to_bl (w AND mask n))) \ n" proof - have "length (takeWhile Not (replicate n False @ ls)) = n + length (takeWhile Not ls)" for ls n by(subst takeWhile_append2) simp+ then show ?thesis unfolding bl_and_mask by (simp add: dropWhile_eq_drop) qed lemma minus_one_word: "(-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" by simp lemma mask_exceed: "n \ LENGTH('a) \ (x::'a::len word) && ~~ mask n = 0" by (simp add: and_not_mask shiftr_eq_0) lemma two_power_strict_part_mono: "strict_part_mono {..LENGTH('a) - 1} (\x. (2 :: 'a :: len word) ^ x)" proof - { fix n have "n < LENGTH('a) \ strict_part_mono {..n} (\x. (2 :: 'a :: len word) ^ x)" proof (induct n) case 0 then show ?case by simp next case (Suc n) from Suc.prems have "2 ^ n < (2 :: 'a :: len word) ^ Suc n" using power_strict_increasing unat_power_lower word_less_nat_alt by fastforce with Suc show ?case by (subst strict_part_mono_by_steps) simp qed } then show ?thesis by simp qed lemma word_shift_by_2: "x * 4 = (x::'a::len word) << 2" by (simp add: shiftl_t2n) (* simp normal form would be "nat (bintrunc (LENGTH('a)) 2) = 2" *) lemma word_len_min_2: "Suc 0 < LENGTH('a) \ unat (2::'a::len word) = 2" by (metis less_trans_Suc n_less_equal_power_2 numeral_2_eq_2 of_nat_numeral unat_of_nat_len) lemma upper_bits_unset_is_l2p: "n < LENGTH('a) \ (\n' \ n. n' < LENGTH('a) \ \ p !! n') = ((p::'a::len word) < 2 ^ n)" apply (cases "Suc 0 < LENGTH('a)") prefer 2 apply (subgoal_tac "LENGTH('a) = 1", auto simp: word_eq_iff)[1] apply (rule iffI) apply (subst mask_eq_iff_w2p [symmetric]) apply (clarsimp simp: word_size) apply (rule word_eqI, rename_tac n') apply (case_tac "n' < n"; simp add: word_size) apply clarify apply (drule bang_is_le) apply (drule_tac y=p in order_le_less_trans, assumption) apply (drule word_power_increasing; simp add: word_len_min_2[simplified]) done lemma less_2p_is_upper_bits_unset: "((p::'a::len word) < 2 ^ n) = (n < LENGTH('a) \ (\n' \ n. n' < LENGTH('a) \ \ p !! n'))" apply (cases "n < LENGTH('a)") apply (simp add: upper_bits_unset_is_l2p) apply (simp add: power_overflow) done lemma le_2p_upper_bits: "\ (p::'a::len word) \ 2^n - 1; n < LENGTH('a) \ \ \n'\n. n' < LENGTH('a) \ \ p !! n'" by (subst upper_bits_unset_is_l2p; simp) lemma le2p_bits_unset: "p \ 2 ^ n - 1 \ \n'\n. n' < LENGTH('a) \ \ (p::'a::len word) !! n'" using upper_bits_unset_is_l2p [where p=p] by (cases "n < LENGTH('a)") auto lemma ucast_less_shiftl_helper: "\ LENGTH('b) + 2 < LENGTH('a); 2 ^ (LENGTH('b) + 2) \ n\ \ (ucast (x :: 'b::len word) << 2) < (n :: 'a::len word)" apply (erule order_less_le_trans[rotated]) using ucast_less[where x=x and 'a='a] apply (simp only: shiftl_t2n field_simps) apply (rule word_less_power_trans2; simp) done lemma word_power_nonzero: "\ (x :: 'a::len word) < 2 ^ (LENGTH('a) - n); n < LENGTH('a); x \ 0 \ \ x * 2 ^ n \ 0" apply (cases "n = 0", simp) apply (simp only: word_neq_0_conv word_less_nat_alt shiftl_t2n mod_0 unat_word_ariths unat_power_lower word_le_nat_alt) apply (subst mod_less) apply (subst mult.commute, erule nat_less_power_trans) apply simp apply simp done lemma less_1_helper: "n \ m \ (n - 1 :: int) < m" by arith lemma div_power_helper: "\ x \ y; y < LENGTH('a) \ \ (2 ^ y - 1) div (2 ^ x :: 'a::len word) = 2 ^ (y - x) - 1" apply (rule word_uint.Rep_eqD) apply (simp only: uint_word_ariths uint_div uint_power_lower) apply (subst mod_pos_pos_trivial, fastforce, fastforce)+ apply (subst mod_pos_pos_trivial) apply (simp add: le_diff_eq uint_2p_alt) apply (rule less_1_helper) apply (rule power_increasing; simp) apply (subst mod_pos_pos_trivial) apply (simp add: uint_2p_alt) apply (rule less_1_helper) apply (rule power_increasing; simp) apply (subst int_div_sub_1; simp add: uint_2p_alt) apply (subst power_0[symmetric]) apply (simp add: uint_2p_alt le_imp_power_dvd power_sub_int) done lemma word_add_power_off: fixes a :: "'a :: len word" assumes ak: "a < k" and kw: "k < 2 ^ (LENGTH('a) - m)" and mw: "m < LENGTH('a)" and off: "off < 2 ^ m" shows "(a * 2 ^ m) + off < k * 2 ^ m" proof (cases "m = 0") case True then show ?thesis using off ak by simp next case False from ak have ak1: "a + 1 \ k" by (rule inc_le) then have "(a + 1) * 2 ^ m \ 0" apply - apply (rule word_power_nonzero) apply (erule order_le_less_trans [OF _ kw]) apply (rule mw) apply (rule less_is_non_zero_p1 [OF ak]) done then have "(a * 2 ^ m) + off < ((a + 1) * 2 ^ m)" using kw mw apply - apply (simp add: distrib_right) apply (rule word_plus_strict_mono_right [OF off]) apply (rule is_aligned_no_overflow'') apply (rule is_aligned_mult_triv2) apply assumption done also have "\ \ k * 2 ^ m" using ak1 mw kw False apply - apply (erule word_mult_le_mono1) apply (simp add: p2_gt_0) apply (simp add: word_less_nat_alt) apply (rule nat_less_power_trans2[simplified]) apply (simp add: word_less_nat_alt) apply simp done finally show ?thesis . qed lemma offset_not_aligned: "\ is_aligned (p::'a::len word) n; i > 0; i < 2 ^ n; n < LENGTH('a)\ \ \ is_aligned (p + of_nat i) n" apply (erule is_aligned_add_not_aligned) unfolding is_aligned_def apply clarsimp apply (subst (asm) unat_of_nat_len) apply (metis order_less_trans unat_lt2p unat_power_lower) apply (metis nat_dvd_not_less) done lemma length_upto_enum_one: fixes x :: "'a :: len word" assumes lt1: "x < y" and lt2: "z < y" and lt3: "x \ z" shows "[x , y .e. z] = [x]" unfolding upto_enum_step_def proof (subst upto_enum_red, subst if_not_P [OF leD [OF lt3]], clarsimp, rule conjI) show "unat ((z - x) div (y - x)) = 0" proof (subst unat_div, rule div_less) have syx: "unat (y - x) = unat y - unat x" by (rule unat_sub [OF order_less_imp_le]) fact moreover have "unat (z - x) = unat z - unat x" by (rule unat_sub) fact ultimately show "unat (z - x) < unat (y - x)" using lt3 apply simp apply (rule diff_less_mono[OF unat_mono, OF lt2]) apply (simp add: word_le_nat_alt[symmetric]) done qed then show "(z - x) div (y - x) * (y - x) = 0" by (metis mult_zero_left unat_0 word_unat.Rep_eqD) qed lemma max_word_mask: shows "(max_word :: 'a::len word) = mask (LENGTH('a))" unfolding mask_def by (metis max_word_eq shiftl_1) lemma is_aligned_alignUp[simp]: "is_aligned (alignUp p n) n" by (simp add: alignUp_def complement_def is_aligned_mask mask_def word_bw_assocs) lemma alignUp_le[simp]: "alignUp p n \ p + 2 ^ n - 1" unfolding alignUp_def by (rule word_and_le2) lemma complement_mask: "complement (2 ^ n - 1) = ~~ mask n" unfolding complement_def mask_def by simp lemma alignUp_idem: fixes a :: "'a::len word" assumes al: "is_aligned a n" and sz: "n < LENGTH('a)" shows "alignUp a n = a" using sz al unfolding alignUp_def apply (simp add: complement_mask) apply (subst x_power_minus_1) apply (subst neg_mask_is_div) apply (simp only: word_arith_nat_div unat_word_ariths) apply (simp only: unat_power_lower) apply (subst power_mod_div) apply (erule is_alignedE) apply simp apply (subst unat_mult_power_lem) apply simp apply (subst unat_sub) apply (subst unat_arith_simps) apply simp apply (simp add: del: unat_1) apply simp done lemma alignUp_not_aligned_eq: fixes a :: "'a :: len word" assumes al: "\ is_aligned a n" and sz: "n < LENGTH('a)" shows "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" proof - have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) fact+ then have um: "unat (a mod 2 ^ n - 1) div 2 ^ n = 0" using sz apply - apply (rule div_less) apply (simp add: unat_minus_one) apply (rule order_less_trans) apply (rule diff_Suc_less) apply (erule contrapos_np) apply (simp add: unat_eq_zero) apply (subst unat_power_lower [symmetric, OF sz]) apply (subst word_less_nat_alt [symmetric]) apply (rule word_mod_less_divisor) apply (simp add: p2_gt_0) done have "a + 2 ^ n - 1 = (a div 2 ^ n) * 2 ^ n + (a mod 2 ^ n) + 2 ^ n - 1" by (simp add: word_mod_div_equality) also have "\ = (a mod 2 ^ n - 1) + (a div 2 ^ n + 1) * 2 ^ n" by (simp add: field_simps) finally show "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" using sz unfolding alignUp_def apply (subst complement_mask) apply (erule ssubst) apply (subst neg_mask_is_div) apply (simp add: word_arith_nat_div) apply (subst unat_word_ariths(1) unat_word_ariths(2))+ apply (subst uno_simps) apply (subst unat_1) apply (subst mod_add_right_eq) apply simp apply (subst power_mod_div) apply (subst div_mult_self1) apply simp apply (subst um) apply simp apply (subst mod_mod_power) apply simp apply (subst word_unat_power, subst Abs_fnat_hom_mult) apply (subst mult_mod_left) apply (subst power_add [symmetric]) apply simp apply (subst Abs_fnat_hom_1) apply (subst Abs_fnat_hom_add) apply (subst word_unat_power, subst Abs_fnat_hom_mult) apply (subst word_unat.Rep_inverse[symmetric], subst Abs_fnat_hom_mult) apply simp done qed lemma alignUp_ge: fixes a :: "'a :: len word" assumes sz: "n < LENGTH('a)" and nowrap: "alignUp a n \ 0" shows "a \ alignUp a n" proof (cases "is_aligned a n") case True then show ?thesis using sz by (subst alignUp_idem, simp_all) next case False have lt0: "unat a div 2 ^ n < 2 ^ (LENGTH('a) - n)" using sz apply - apply (subst td_gal_lt [symmetric]) apply simp apply (simp add: power_add [symmetric]) done have"2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using sz apply - apply (rule nat_le_power_trans) apply simp apply (rule Suc_leI [OF lt0]) apply simp done moreover have "2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using nowrap sz apply - apply (erule contrapos_nn) apply (subst alignUp_not_aligned_eq [OF False sz]) apply (subst unat_arith_simps) apply (subst unat_word_ariths) apply (subst unat_word_ariths) apply simp apply (subst mult_mod_left) apply (simp add: unat_div field_simps power_add[symmetric] mod_mod_power min.absorb2) done ultimately have lt: "2 ^ n * (unat a div 2 ^ n + 1) < 2 ^ LENGTH('a)" by simp have "a = a div 2 ^ n * 2 ^ n + a mod 2 ^ n" by (rule word_mod_div_equality [symmetric]) also have "\ < (a div 2 ^ n + 1) * 2 ^ n" using sz lt apply (simp add: field_simps) apply (rule word_add_less_mono1) apply (rule word_mod_less_divisor) apply (simp add: word_less_nat_alt) apply (subst unat_word_ariths) apply (simp add: unat_div) done also have "\ = alignUp a n" by (rule alignUp_not_aligned_eq [symmetric]) fact+ finally show ?thesis by (rule order_less_imp_le) qed lemma alignUp_le_greater_al: fixes x :: "'a :: len word" assumes le: "a \ x" and sz: "n < LENGTH('a)" and al: "is_aligned x n" shows "alignUp a n \ x" proof (cases "is_aligned a n") case True then show ?thesis using sz le by (simp add: alignUp_idem) next case False then have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) from al obtain k where xk: "x = 2 ^ n * of_nat k" and kv: "k < 2 ^ (LENGTH('a) - n)" by (auto elim!: is_alignedE) then have kn: "unat (of_nat k :: 'a word) * unat ((2::'a word) ^ n) < 2 ^ LENGTH('a)" using sz apply (subst unat_of_nat_eq) apply (erule order_less_le_trans) apply simp apply (subst mult.commute) apply simp apply (rule nat_less_power_trans) apply simp apply simp done have au: "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" by (rule alignUp_not_aligned_eq) fact+ also have "\ \ of_nat k * 2 ^ n" proof (rule word_mult_le_mono1 [OF inc_le _ kn]) show "a div 2 ^ n < of_nat k" using kv xk le sz anz by (simp add: alignUp_div_helper) show "(0:: 'a word) < 2 ^ n" using sz by (simp add: p2_gt_0 sz) qed finally show ?thesis using xk by (simp add: field_simps) qed lemma alignUp_is_aligned_nz: fixes a :: "'a :: len word" assumes al: "is_aligned x n" and sz: "n < LENGTH('a)" and ax: "a \ x" and az: "a \ 0" shows "alignUp (a::'a :: len word) n \ 0" proof (cases "is_aligned a n") case True then have "alignUp a n = a" using sz by (simp add: alignUp_idem) then show ?thesis using az by simp next case False then have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) { assume asm: "alignUp a n = 0" have lt0: "unat a div 2 ^ n < 2 ^ (LENGTH('a) - n)" using sz apply - apply (subst td_gal_lt [symmetric]) apply simp apply (simp add: power_add [symmetric]) done have leq: "2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using sz apply - apply (rule nat_le_power_trans) apply simp apply (rule Suc_leI [OF lt0]) apply simp done from al obtain k where kv: "k < 2 ^ (LENGTH('a) - n)" and xk: "x = 2 ^ n * of_nat k" by (auto elim!: is_alignedE) then have "a div 2 ^ n < of_nat k" using ax sz anz by (rule alignUp_div_helper) then have r: "unat a div 2 ^ n < k" using sz apply (simp add: unat_div word_less_nat_alt) apply (subst (asm) unat_of_nat) apply (subst (asm) mod_less) apply (rule order_less_le_trans [OF kv]) apply simp+ done have "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" by (rule alignUp_not_aligned_eq) fact+ then have "\ = 0" using asm by simp then have "2 ^ LENGTH('a) dvd 2 ^ n * (unat a div 2 ^ n + 1)" using sz by (simp add: unat_arith_simps ac_simps) (simp add: unat_word_ariths mod_simps mod_eq_0_iff_dvd) with leq have "2 ^ n * (unat a div 2 ^ n + 1) = 2 ^ LENGTH('a)" by (force elim!: le_SucE) then have "unat a div 2 ^ n = 2 ^ LENGTH('a) div 2 ^ n - 1" by (metis (no_types, hide_lams) Groups.add_ac(2) add.right_neutral add_diff_cancel_left' div_le_dividend div_mult_self4 gr_implies_not0 le_neq_implies_less power_eq_0_iff unat_def zero_neq_numeral) then have "unat a div 2 ^ n = 2 ^ (LENGTH('a) - n) - 1" using sz by (simp add: power_sub) then have "2 ^ (LENGTH('a) - n) - 1 < k" using r by simp then have False using kv by simp } then show ?thesis by clarsimp qed lemma alignUp_ar_helper: fixes a :: "'a :: len word" assumes al: "is_aligned x n" and sz: "n < LENGTH('a)" and sub: "{x..x + 2 ^ n - 1} \ {a..b}" and anz: "a \ 0" shows "a \ alignUp a n \ alignUp a n + 2 ^ n - 1 \ b" proof from al have xl: "x \ x + 2 ^ n - 1" by (simp add: is_aligned_no_overflow) from xl sub have ax: "a \ x" by (clarsimp elim!: range_subset_lower [where x = x]) show "a \ alignUp a n" proof (rule alignUp_ge) show "alignUp a n \ 0" using al sz ax anz by (rule alignUp_is_aligned_nz) qed fact+ show "alignUp a n + 2 ^ n - 1 \ b" proof (rule order_trans) from xl show tp: "x + 2 ^ n - 1 \ b" using sub by (clarsimp elim!: range_subset_upper [where x = x]) from ax have "alignUp a n \ x" by (rule alignUp_le_greater_al) fact+ then have "alignUp a n + (2 ^ n - 1) \ x + (2 ^ n - 1)" using xl al is_aligned_no_overflow' olen_add_eqv word_plus_mcs_3 by blast then show "alignUp a n + 2 ^ n - 1 \ x + 2 ^ n - 1" by (simp add: field_simps) qed qed lemma alignUp_def2: "alignUp a sz = a + 2 ^ sz - 1 && ~~ mask sz" unfolding alignUp_def[unfolded complement_def] by (simp add:mask_def[symmetric,unfolded shiftl_t2n,simplified]) lemma mask_out_first_mask_some: "\ x && ~~ mask n = y; n \ m \ \ x && ~~ mask m = y && ~~ mask m" apply (rule word_eqI, rename_tac n') apply (drule_tac x=n' in word_eqD) apply (auto simp: word_ops_nth_size word_size) done lemma gap_between_aligned: "\a < (b :: 'a ::len word); is_aligned a n; is_aligned b n; n < LENGTH('a) \ \ a + (2^n - 1) < b" apply (rule ccontr,simp add:not_less) apply (drule le_shiftr[where n = n]) apply (simp add: aligned_shift') apply (case_tac "b >> n = a >> n") apply (drule arg_cong[where f = "\x. x< alignUp (w + a) us = w + alignUp a us" apply (clarsimp simp:alignUp_def2 add.assoc) apply (simp add: mask_out_add_aligned field_simps) done lemma mask_lower_twice: "n \ m \ (x && ~~ mask n) && ~~ mask m = x && ~~ mask m" apply (rule word_eqI) apply (simp add: word_size word_ops_nth_size) apply safe apply simp done lemma mask_lower_twice2: "(a && ~~ mask n) && ~~ mask m = a && ~~ mask (max n m)" by (rule word_eqI, simp add: neg_mask_bang conj_comms) lemma ucast_and_neg_mask: "ucast (x && ~~ mask n) = ucast x && ~~ mask n" apply (rule word_eqI) apply (simp add: word_size neg_mask_bang nth_ucast) apply (auto simp add: test_bit_bl word_size) done lemma ucast_and_mask: "ucast (x && mask n) = ucast x && mask n" apply (rule word_eqI) apply (simp add: nth_ucast) apply (auto simp add: test_bit_bl word_size) done lemma ucast_mask_drop: "LENGTH('a :: len) \ n \ (ucast (x && mask n) :: 'a word) = ucast x" apply (rule word_eqI) apply (simp add: nth_ucast word_size) apply safe apply (simp add: test_bit_bl word_size) done lemma alignUp_distance: "alignUp (q :: 'a :: len word) sz - q \ mask sz" apply (case_tac "LENGTH('a) \ sz") apply (simp add:alignUp_def2 mask_def power_overflow) apply (case_tac "is_aligned q sz") apply (clarsimp simp:alignUp_def2 p_assoc_help) apply (subst mask_out_add_aligned[symmetric],simp)+ apply (simp add:mask_lower_twice word_and_le2) apply (simp add:and_not_mask) apply (subst le_mask_iff[THEN iffD1]) apply (simp add:mask_def) apply simp apply (clarsimp simp:alignUp_def3) apply (subgoal_tac "2 ^ sz - (q - (q - 1 && ~~ mask sz)) \ 2 ^ sz - 1") apply (simp add:field_simps mask_def) apply (rule word_sub_mono) apply simp apply (rule ccontr) apply (clarsimp simp:not_le) apply (drule eq_refl) apply (drule order_trans[OF _ word_and_le2]) apply (subgoal_tac "q \ 0") apply (drule minus_one_helper[rotated]) apply simp apply simp apply (fastforce) apply (simp add: word_sub_le_iff) apply (subgoal_tac "q - 1 && ~~ mask sz = (q - 1) - (q - 1 && mask sz)") apply simp apply (rule order_trans) apply (rule word_add_le_mono2) apply (rule word_and_le1) apply (subst unat_plus_simple[THEN iffD1,symmetric]) apply (simp add:not_le mask_def) apply (rule word_sub_1_le) apply simp apply (rule unat_lt2p) apply (simp add:mask_def) apply (simp add:mask_out_sub_mask) apply (rule word_sub_1_le) apply simp done lemma is_aligned_diff_neg_mask: "is_aligned p sz \ (p - q && ~~ mask sz) = (p - ((alignUp q sz) && ~~ mask sz))" apply (clarsimp simp only:word_and_le2 diff_conv_add_uminus) apply (subst mask_out_add_aligned[symmetric]) apply simp+ apply (rule sum_to_zero) apply (subst add.commute) apply (subst mask_out_add_aligned) apply (simp add:is_aligned_neg_mask) apply simp apply (subst and_not_mask[where w = "(alignUp q sz && ~~ mask sz) - q "]) apply (subst le_mask_iff[THEN iffD1]) apply (simp add:is_aligned_neg_mask_eq) apply (rule alignUp_distance) apply simp done lemma map_bits_rev_to_bl: "map ((!!) x) [0.. LENGTH('a)" shows "x = ucast y \ ucast x = y" using assms by (simp add: is_down ucast_id ucast_ucast_a) lemma le_ucast_ucast_le: fixes x :: "'a::len word" and y :: "'b::len word" assumes "LENGTH('b) \ LENGTH('a)" shows "x \ ucast y \ ucast x \ y" using assms apply (simp add: word_le_nat_alt unat_ucast_up_simp[where x=y]) apply (simp add: unat_ucast) by (rule le_trans; fastforce) lemma less_ucast_ucast_less: fixes x :: "'a::len word" and y :: "'b::len word" assumes "LENGTH('b) \ LENGTH('a)" shows "x < ucast y \ ucast x < y" using assms apply (simp add: word_less_nat_alt unat_ucast_up_simp[where x=y]) apply (simp add: unat_ucast) by (rule le_less_trans; fastforce) lemma ucast_le_ucast: fixes x :: "'a::len word" and y :: "'a::len word" assumes "LENGTH('a) \ LENGTH('b)" shows "(ucast x \ (ucast y::'b::len word)) = (x \ y)" using assms apply (simp add: word_le_nat_alt unat_ucast) apply (subst mod_less) apply(rule less_le_trans[OF unat_lt2p], simp) apply (subst mod_less) apply(rule less_le_trans[OF unat_lt2p], simp) apply simp done lemma ucast_le_ucast_eq: fixes x y :: "'a::len word" assumes x: "x < 2 ^ n" assumes y: "y < 2 ^ n" assumes n: "n = LENGTH('b::len)" shows "(UCAST('a \ 'b) x \ UCAST('a \ 'b) y) = (x \ y)" apply (rule iffI) apply (cases "LENGTH('b) < LENGTH('a)") apply (subst less_mask_eq[OF x, symmetric]) apply (subst less_mask_eq[OF y, symmetric]) apply (unfold n) apply (subst ucast_ucast_mask[symmetric])+ apply (simp add: ucast_le_ucast)+ apply (erule ucast_mono_le[OF _ y[unfolded n]]) done (* High bits w.r.t. mask operations. *) lemma and_neg_mask_eq_iff_not_mask_le: "w && ~~ mask n = ~~ mask n \ ~~ mask n \ w" by (metis (full_types) dual_order.antisym neg_mask_mono_le word_and_le1 word_and_le2 word_bool_alg.conj_absorb) lemma le_mask_high_bits: shows "w \ mask n \ (\ i \ {n ..< size w}. \ w !! i)" by (auto simp: word_size and_mask_eq_iff_le_mask[symmetric] word_eq_iff) lemma neg_mask_le_high_bits: shows "~~ mask n \ w \ (\ i \ {n ..< size w}. w !! i)" by (auto simp: word_size and_neg_mask_eq_iff_not_mask_le[symmetric] word_eq_iff neg_mask_bang) lemma word_le_not_less: "((b::'a::len word) \ a) = (\(a < b))" by fastforce lemma ucast_or_distrib: fixes x :: "'a::len word" fixes y :: "'a::len word" shows "(ucast (x || y) :: ('b::len) word) = ucast x || ucast y" unfolding ucast_def Word.bitOR_word.abs_eq uint_or by blast lemma shiftr_less: "(w::'a::len word) < k \ w >> n < k" apply (simp add: word_less_nat_alt shiftr_div_2n') apply (blast intro: div_le_dividend le_less_trans) done lemma word_and_notzeroD: "w && w' \ 0 \ w \ 0 \ w' \ 0" by auto lemma word_clz_max: "word_clz w \ size (w::'a::len word)" unfolding word_clz_def apply (simp add: word_size) apply (rule_tac y="length (to_bl w)" in order_trans) apply (rule List.length_takeWhile_le) apply simp done lemma word_clz_nonzero_max: fixes w :: "'a::len word" assumes nz: "w \ 0" shows "word_clz w < size (w::'a::len word)" proof - { assume a: "word_clz w = size (w::'a::len word)" hence "length (takeWhile Not (to_bl w)) = length (to_bl w)" by (simp add: word_clz_def word_size) hence allj: "\j\set(to_bl w). \ j" using takeWhile_take_has_property[where n="length (to_bl w)" and xs="to_bl w" and P=Not] by simp hence "to_bl w = replicate (length (to_bl w)) False" by (fastforce intro!: list_of_false) hence "w = 0" apply simp apply (subst (asm) to_bl_0[symmetric]) apply (drule Word.word_bl.Rep_eqD, assumption) done with nz have False by simp } thus ?thesis using word_clz_max by (fastforce intro: le_neq_trans) qed lemma unat_add_lem': "(unat x + unat y < 2 ^ LENGTH('a)) \ (unat (x + y :: 'a :: len word) = unat x + unat y)" by (subst unat_add_lem[symmetric], assumption) lemma from_bool_eq_if': "((if P then 1 else 0) = from_bool Q) = (P = Q)" by (simp add: case_bool_If from_bool_def split: if_split) lemma word_exists_nth: "(w::'a::len word) \ 0 \ \i. w !! i" using word_log2_nth_same by blast lemma shiftr_le_0: "unat (w::'a::len word) < 2 ^ n \ w >> n = (0::'a::len word)" by (rule word_unat.Rep_eqD) (simp add: shiftr_div_2n') lemma of_nat_shiftl: "(of_nat x << n) = (of_nat (x * 2 ^ n) :: ('a::len) word)" proof - have "(of_nat x::'a word) << n = of_nat (2 ^ n) * of_nat x" using shiftl_t2n by (metis word_unat_power) thus ?thesis by simp qed lemma shiftl_1_not_0: "n < LENGTH('a) \ (1::'a::len word) << n \ 0" by (simp add: shiftl_t2n) lemma max_word_not_0[simp]: "max_word \ 0" by (simp add: max_word_minus) lemma ucast_zero_is_aligned: "UCAST('a::len \ 'b::len) w = 0 \ n \ LENGTH('b) \ is_aligned w n" by (clarsimp simp: is_aligned_mask word_eq_iff word_size nth_ucast) lemma unat_ucast_eq_unat_and_mask: "unat (UCAST('b::len \ 'a::len) w) = unat (w && mask LENGTH('a))" proof - have "unat (UCAST('b \ 'a) w) = unat (UCAST('a \ 'b) (UCAST('b \ 'a) w))" by (cases "LENGTH('a) < LENGTH('b)"; simp add: is_down ucast_ucast_a unat_ucast_up_simp) thus ?thesis using ucast_ucast_mask by simp qed lemma unat_max_word_pos[simp]: "0 < unat max_word" by (auto simp: unat_gt_0) (* Miscellaneous conditional injectivity rules. *) lemma mult_pow2_inj: assumes ws: "m + n \ LENGTH('a)" assumes le: "x \ mask m" "y \ mask m" assumes eq: "x * 2^n = y * (2^n::'a::len word)" shows "x = y" proof (cases n) case 0 thus ?thesis using eq by simp next case (Suc n') have m_lt: "m < LENGTH('a)" using Suc ws by simp have xylt: "x < 2^m" "y < 2^m" using le m_lt unfolding mask_2pm1 by auto have lenm: "n \ LENGTH('a) - m" using ws by simp show ?thesis using eq xylt apply (fold shiftl_t2n[where n=n, simplified mult.commute]) apply (simp only: word_bl.Rep_inject[symmetric] bl_shiftl) apply (erule ssubst[OF less_is_drop_replicate])+ apply (clarsimp elim!: drop_eq_mono[OF lenm]) done qed lemma word_of_nat_inj: assumes bounded: "x < 2 ^ LENGTH('a)" "y < 2 ^ LENGTH('a)" assumes of_nats: "of_nat x = (of_nat y :: 'a::len word)" shows "x = y" by (rule contrapos_pp[OF of_nats]; cases "x < y"; cases "y < x") (auto dest: bounded[THEN of_nat_mono_maybe]) (* Sign extension from bit n. *) lemma sign_extend_bitwise_if: "i < size w \ sign_extend e w !! i \ (if i < e then w !! i else w !! e)" by (simp add: sign_extend_def neg_mask_bang word_size) lemma sign_extend_bitwise_disj: "i < size w \ sign_extend e w !! i \ i \ e \ w !! i \ e \ i \ w !! e" by (auto simp: sign_extend_bitwise_if) lemma sign_extend_bitwise_cases: "i < size w \ sign_extend e w !! i \ (i \ e \ w !! i) \ (e \ i \ w !! e)" by (auto simp: sign_extend_bitwise_if) lemmas sign_extend_bitwise_if' = sign_extend_bitwise_if[simplified word_size] lemmas sign_extend_bitwise_disj' = sign_extend_bitwise_disj[simplified word_size] lemmas sign_extend_bitwise_cases' = sign_extend_bitwise_cases[simplified word_size] (* Often, it is easier to reason about an operation which does not overwrite the bit which determines which mask operation to apply. *) lemma sign_extend_def': "sign_extend n w = (if w !! n then w || ~~ mask (Suc n) else w && mask (Suc n))" by (rule word_eqI[rule_format]) (auto simp: sign_extend_bitwise_if' word_size word_ops_nth_size dest: less_antisym) lemma sign_extended_sign_extend: "sign_extended n (sign_extend n w)" by (clarsimp simp: sign_extended_def word_size sign_extend_bitwise_if') lemma sign_extended_iff_sign_extend: "sign_extended n w \ sign_extend n w = w" apply (rule iffI) apply (rule word_eqI[rule_format], rename_tac i) apply (case_tac "n < i"; simp add: sign_extended_def word_size sign_extend_bitwise_if') apply (erule subst, rule sign_extended_sign_extend) done lemma sign_extended_weaken: "sign_extended n w \ n \ m \ sign_extended m w" unfolding sign_extended_def by (cases "n < m") auto lemma sign_extend_sign_extend_eq: "sign_extend m (sign_extend n w) = sign_extend (min m n) w" by (cases "m < n") (auto intro!: word_eqI simp: word_size sign_extend_bitwise_cases') lemma sign_extended_high_bits: "\ sign_extended e p; j < size p; e \ i; i < j \ \ p !! i = p !! j" by (drule (1) sign_extended_weaken; simp add: sign_extended_def) lemma sign_extend_eq: "w && mask (Suc n) = v && mask (Suc n) \ sign_extend n w = sign_extend n v" by (rule word_eqI, fastforce dest: word_eqD simp: sign_extend_bitwise_if' word_size) lemma sign_extended_add: assumes p: "is_aligned p n" assumes f: "f < 2 ^ n" assumes e: "n \ e" assumes "sign_extended e p" shows "sign_extended e (p + f)" proof (cases "e < size p") case True note and_or = is_aligned_add_or[OF p f] have "\ f !! e" using True e less_2p_is_upper_bits_unset[THEN iffD1, OF f] by (fastforce simp: word_size) hence i: "(p + f) !! e = p !! e" by (simp add: and_or) have fm: "f && mask e = f" by (fastforce intro: subst[where P="\f. f && mask e = f", OF less_mask_eq[OF f]] simp: mask_twice e) show ?thesis using assms apply (simp add: sign_extended_iff_sign_extend sign_extend_def i) apply (simp add: and_or word_bw_comms[of p f]) apply (clarsimp simp: word_ao_dist fm word_bw_assocs split: if_splits) done next case False thus ?thesis by (simp add: sign_extended_def word_size) qed lemma sign_extended_neq_mask: "\sign_extended n ptr; m \ n\ \ sign_extended n (ptr && ~~ mask m)" by (fastforce simp: sign_extended_def word_size neg_mask_bang) (* Uints *) lemma uints_mono_iff: "uints l \ uints m \ l \ m" using power_increasing_iff[of "2::int" l m] apply (auto simp: uints_num subset_iff simp del: power_increasing_iff) by (meson less_irrefl not_less zle2p) lemmas uints_monoI = uints_mono_iff[THEN iffD2] lemma Bit_in_uints_Suc: "w BIT c \ uints (Suc m)" if "w \ uints m" using that by (auto simp: uints_num Bit_def) lemma Bit_in_uintsI: "w BIT c \ uints m" if "w \ uints (m - 1)" "m > 0" using Bit_in_uints_Suc[OF that(1)] that(2) by auto lemma bin_cat_in_uintsI: "bin_cat a n b \ uints m" if "a \ uints l" "m \ l + n" using that proof (induction n arbitrary: b m) case 0 then have "uints l \ uints m" by (intro uints_monoI) auto then show ?case using 0 by auto next case (Suc n) then show ?case by (auto intro!: Bit_in_uintsI) qed lemma bin_cat_cong: "bin_cat a n b = bin_cat c m d" if "n = m" "a = c" "bintrunc m b = bintrunc m d" using that(3) unfolding that(1,2) proof (induction m arbitrary: b d) case (Suc m) show ?case using Suc.prems by (auto intro: Suc.IH) qed simp lemma bin_cat_eqD1: "bin_cat a n b = bin_cat c n d \ a = c" by (induction n arbitrary: b d) auto lemma bin_cat_eqD2: "bin_cat a n b = bin_cat c n d \ bintrunc n b = bintrunc n d" by (induction n arbitrary: b d) auto lemma bin_cat_inj: "(bin_cat a n b) = bin_cat c n d \ a = c \ bintrunc n b = bintrunc n d" by (auto intro: bin_cat_cong bin_cat_eqD1 bin_cat_eqD2) lemma word_of_int_bin_cat_eq_iff: "(word_of_int (bin_cat (uint a) LENGTH('b) (uint b))::'c::len0 word) = word_of_int (bin_cat (uint c) LENGTH('b) (uint d)) \ b = d \ a = c" if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" for a::"'a::len0 word" and b::"'b::len0 word" by (subst word_uint.Abs_inject) (auto simp: bin_cat_inj intro!: that bin_cat_in_uintsI) lemma word_cat_inj: "(word_cat a b::'c::len0 word) = word_cat c d \ a = c \ b = d" if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" for a::"'a::len0 word" and b::"'b::len0 word" by (auto simp: word_cat_def word_uint.Abs_inject word_of_int_bin_cat_eq_iff that) lemma p2_eq_1: "2 ^ n = (1::'a::len word) \ n = 0" proof - have "2 ^ n = (1::'a word) \ n = 0" by (metis One_nat_def not_less one_less_numeral_iff p2_eq_0 p2_gt_0 power_0 power_0 power_inject_exp semiring_norm(76) unat_power_lower zero_neq_one) then show ?thesis by auto qed (* usually: x,y = (len_of TYPE ('a)) *) lemma bitmagic_zeroLast_leq_or1Last: "(a::('a::len) word) AND (mask len << x - len) \ a OR mask (y - len)" by (meson le_word_or2 order_trans word_and_le2) lemma zero_base_lsb_imp_set_eq_as_bit_operation: fixes base ::"'a::len word" assumes valid_prefix: "mask (LENGTH('a) - len) AND base = 0" shows "(base = NOT mask (LENGTH('a) - len) AND a) \ (a \ {base .. base OR mask (LENGTH('a) - len)})" proof have helper3: "x OR y = x OR y AND NOT x" for x y ::"'a::len word" by (simp add: word_oa_dist2) from assms show "base = NOT mask (LENGTH('a) - len) AND a \ a \ {base..base OR mask (LENGTH('a) - len)}" apply(simp add: word_and_le1) apply(metis helper3 le_word_or2 word_bw_comms(1) word_bw_comms(2)) done next assume "a \ {base..base OR mask (LENGTH('a) - len)}" hence a: "base \ a \ a \ base OR mask (LENGTH('a) - len)" by simp show "base = NOT mask (LENGTH('a) - len) AND a" proof - have f2: "\x\<^sub>0. base AND NOT mask x\<^sub>0 \ a AND NOT mask x\<^sub>0" using a neg_mask_mono_le by blast have f3: "\x\<^sub>0. a AND NOT mask x\<^sub>0 \ (base OR mask (LENGTH('a) - len)) AND NOT mask x\<^sub>0" using a neg_mask_mono_le by blast have f4: "base = base AND NOT mask (LENGTH('a) - len)" using valid_prefix by (metis mask_eq_0_eq_x word_bw_comms(1)) hence f5: "\x\<^sub>6. (base OR x\<^sub>6) AND NOT mask (LENGTH('a) - len) = base OR x\<^sub>6 AND NOT mask (LENGTH('a) - len)" using word_ao_dist by (metis) have f6: "\x\<^sub>2 x\<^sub>3. a AND NOT mask x\<^sub>2 \ x\<^sub>3 \ \ (base OR mask (LENGTH('a) - len)) AND NOT mask x\<^sub>2 \ x\<^sub>3" using f3 dual_order.trans by auto have "base = (base OR mask (LENGTH('a) - len)) AND NOT mask (LENGTH('a) - len)" using f5 by auto hence "base = a AND NOT mask (LENGTH('a) - len)" using f2 f4 f6 by (metis eq_iff) thus "base = NOT mask (LENGTH('a) - len) AND a" by (metis word_bw_comms(1)) qed qed lemma unat_minus_one_word: "unat (-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" by (subst minus_one_word) (subst unat_sub_if', clarsimp) lemma of_nat_eq_signed_scast: "(of_nat x = (y :: ('a::len) signed word)) = (of_nat x = (scast y :: 'a word))" by (metis scast_of_nat scast_scast_id(2)) lemma word_ctz_le: "word_ctz (w :: ('a::len word)) \ LENGTH('a)" apply (clarsimp simp: word_ctz_def) apply (rule nat_le_Suc_less_imp[where y="LENGTH('a) + 1" , simplified]) apply (rule order_le_less_trans[OF List.length_takeWhile_le]) apply simp done lemma word_ctz_less: "w \ 0 \ word_ctz (w :: ('a::len word)) < LENGTH('a)" apply (clarsimp simp: word_ctz_def eq_zero_set_bl) apply (rule order_less_le_trans[OF length_takeWhile_less]) apply fastforce+ done lemma word_ctz_not_minus_1: "1 < LENGTH('a) \ of_nat (word_ctz (w :: ('a :: len) word)) \ (- 1 :: ('a::len) word)" apply (cut_tac w=w in word_ctz_le) apply (subst word_unat.Rep_inject[symmetric]) apply (subst unat_of_nat_eq) apply (erule order_le_less_trans, fastforce) apply (subst unat_minus_one_word) apply (rule less_imp_neq) apply (erule order_le_less_trans) apply (subst less_eq_Suc_le) apply (subst le_diff_conv2, fastforce) apply (clarsimp simp: le_diff_conv2 less_eq_Suc_le[symmetric] suc_le_pow_2) done lemmas word_ctz_not_minus_1_32 = word_ctz_not_minus_1[where 'a=32, simplified] end